Skip to content

Commit c7d24e3

Browse files
CohenCyrilgares
authored andcommitted
defining pullbacks
1 parent 5cb5e8a commit c7d24e3

File tree

3 files changed

+177
-19
lines changed

3 files changed

+177
-19
lines changed

_CoqProject

Lines changed: 4 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,10 @@
11
structures.v
2+
theories/cat.v
3+
theories/algebra.v
24
-arg -w -arg -elpi.accumulate-syntax
35
-arg -w -arg +elpi.typecheck
46
-Q . HB
57

68
-R tests HB.tests
7-
-R examples HB.examples
9+
-R examples HB.examples
10+
-R theories HB

theories/algebra.v

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,7 @@
11
Require Import ssreflect ssrfun.
2+
Unset Universe Checking.
23
From HB Require Import structures cat.
4+
Set Universe Checking.
35

46
Set Implicit Arguments.
57
Unset Strict Implicit.

theories/cat.v

Lines changed: 171 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -317,9 +317,10 @@ HB.instance Definition _ {C D : cat} (c : C) :=
317317

318318
(* opposite category *)
319319
Definition catop (C : U) : U := C.
320-
Notation "C ^op" := (catop C) (at level 10, format "C ^op") : cat_scope.
320+
Notation "C ^op" := (catop C)
321+
(at level 2, format "C ^op") : cat_scope.
321322
HB.instance Definition _ (C : quiver) :=
322-
IsQuiver.Build (C^op) (fun a b => hom b a).
323+
IsQuiver.Build C^op (fun a b => hom b a).
323324
HB.instance Definition _ (C : precat) :=
324325
Quiver_IsPreCat.Build (C^op) (fun=> idmap) (fun _ _ _ f g => g \; f).
325326
HB.instance Definition _ (C : cat) := PreCat_IsCat.Build (C^op)
@@ -737,23 +738,23 @@ Notation "F `/ b" := (F `/` cst unit b)
737738
(at level 40, b at level 40, format "F `/ b") : cat_scope.
738739
Notation "a / b" := (cst unit a `/ b) : cat_scope.
739740

740-
(* (* Not working yet *) *)
741-
(* HB.mixin Record IsInitial {C : quiver} (i : C) := { *)
742-
(* to : forall c, i ~> c; *)
743-
(* to_unique : forall c (f : i ~> c), f = to c *)
744-
(* }. *)
745-
(* #[short(type="initial")] *)
746-
(* HB.structure Definition Initial {C : quiver} := {i of IsInitial C i}. *)
741+
Definition obj (C : quiver) : Type := C.
742+
HB.mixin Record IsInitial {C : quiver} (i : obj C) := {
743+
to : forall c, i ~> c;
744+
to_unique : forall c (f : i ~> c), f = to c
745+
}.
746+
#[short(type="initial")]
747+
HB.structure Definition Initial {C : quiver} := {i of IsInitial C i}.
747748

748-
(* HB.mixin Record IsTerminal {C : quiver} (t : C) := { *)
749-
(* from : forall c, c ~> t; *)
750-
(* from_unique : forall c (f : c ~> t), f = from c *)
751-
(* }. *)
752-
(* #[short(type="terminal")] *)
753-
(* HB.structure Definition Terminal {C : quiver} := {t of IsTerminal C t}. *)
754-
(* #[short(type="universal")] *)
755-
(* HB.structure Definition Universal {C : quiver} := *)
756-
(* {u of Initial C u & Terminal C u}. *)
749+
HB.mixin Record IsTerminal {C : quiver} (t : obj C) := {
750+
from : forall c, c ~> t;
751+
from_unique : forall c (f : c ~> t), f = from c
752+
}.
753+
#[short(type="terminal")]
754+
HB.structure Definition Terminal {C : quiver} := {t of IsTerminal C t}.
755+
#[short(type="universal")]
756+
HB.structure Definition Universal {C : quiver} :=
757+
{u of Initial C u & Terminal C u}.
757758

758759
(* Definition hom' {C : precat} (a b : C) := a ~> b. *)
759760
(* (* Bug *) *)
@@ -802,6 +803,7 @@ Arguments L_ {_ _}.
802803
Arguments phi {D C s} {c d}.
803804
Arguments psy {D C s} {c d}.
804805

806+
805807
HB.mixin Record PreCat_IsMonoidal C of PreCat C := {
806808
onec : C;
807809
prod : (C * C)%type ~>_precat C;
@@ -882,3 +884,154 @@ HB.structure Definition Monoidal : Set :=
882884
Set Universe Checking.
883885

884886

887+
Record cospan (Q : quiver) (A B : Q) := Cospan {
888+
top : Q;
889+
left2top : A ~> top;
890+
right2top : B ~> top
891+
}.
892+
893+
Section cospans.
894+
Variables (Q : precat) (A B : Q).
895+
Record cospan_map (c c' : cospan A B) := CospanMap {
896+
top_map : top c ~> top c';
897+
left2top_map : left2top c \; top_map = left2top c';
898+
right2top_map : right2top c \; top_map = right2top c';
899+
}.
900+
HB.instance Definition _ := IsQuiver.Build (cospan A B) cospan_map.
901+
902+
Lemma cospan_mapP (c c' : cospan A B) (f g : c ~> c') :
903+
top_map f = top_map g <-> f = g.
904+
Admitted.
905+
End cospans.
906+
907+
Section cospans_in_cat.
908+
Variables (Q : cat) (A B : Q).
909+
Definition cospan_idmap (c : cospan A B) :=
910+
@CospanMap Q A B c c idmap (compo1 _) (compo1 _).
911+
Program Definition cospan_comp (c1 c2 c3 : cospan A B)
912+
(f12 : c1 ~> c2) (f23 : c2 ~> c3) :=
913+
@CospanMap Q A B c1 c3 (top_map f12 \; top_map f23) _ _.
914+
Next Obligation. by rewrite compoA !left2top_map. Qed.
915+
Next Obligation. by rewrite compoA !right2top_map. Qed.
916+
HB.instance Definition _ := IsPreCat.Build (cospan A B) cospan_idmap cospan_comp.
917+
918+
Lemma cospan_are_cats : PreCat_IsCat (cospan A B).
919+
Proof.
920+
constructor=> [a b f|a b f|a b c d f g h].
921+
- by apply/cospan_mapP => /=; rewrite comp1o.
922+
- by apply/cospan_mapP => /=; rewrite compo1.
923+
- by apply/cospan_mapP => /=; rewrite compoA.
924+
Qed.
925+
HB.instance Definition _ := cospan_are_cats.
926+
End cospans_in_cat.
927+
928+
Section spans_and_co.
929+
Variables (Q : quiver) (A B : Q).
930+
Definition span := @cospan Q^op A B.
931+
Definition bot (s : span) : Q := top s.
932+
Definition bot2left (s : span) : bot s ~>_Q A := left2top s.
933+
Definition bot2right (s : span) : bot s ~>_Q B := right2top s.
934+
Definition Span (C : Q) (f : C ~> A) (g : C ~> B) : span :=
935+
@Cospan Q^op A B C f g.
936+
End spans_and_co.
937+
938+
HB.instance Definition _ (Q : precat) (A B : Q) :=
939+
Quiver.copy (span A B) (@cospan Q^op A B)^op.
940+
HB.instance Definition _ (Q : cat) (A B : Q) :=
941+
Cat.copy (span A B) (@cospan Q^op A B)^op.
942+
943+
Section bot_map.
944+
Variables (C : cat) (A B : C) (s s' : span A B) (f : s ~> s').
945+
Definition bot_map : bot s ~> bot s' := top_map f.
946+
Lemma bot2left_map : bot_map \; bot2left s' = bot2left s.
947+
Proof. exact: left2top_map f. Qed.
948+
Lemma bot2right_map : bot_map \; bot2right s' = bot2right s.
949+
Proof. exact: right2top_map f. Qed.
950+
End bot_map.
951+
952+
HB.mixin Record isPrePullback (Q : precat) (A B : Q)
953+
(c : cospan A B) (s : span A B) := {
954+
is_square : bot2left s \; left2top c = bot2right s \; right2top c;
955+
}.
956+
#[short(type=prepullback)]
957+
HB.structure Definition PrePullback Q A B (c : cospan A B) :=
958+
{s of isPrePullback Q A B c s}.
959+
Arguments prepullback {Q A B} c.
960+
961+
Section prepullback.
962+
Variables (Q : precat) (A B : Q) (c : cospan A B).
963+
HB.instance Definition _ :=
964+
IsQuiver.Build (prepullback c)
965+
(fun a b => (a : span A B) ~>_(span A B) (b : span A B)).
966+
Lemma eq_prepullbackP (p q : prepullback c) :
967+
p = q :> span _ _ <-> p = q.
968+
Admitted.
969+
End prepullback.
970+
Section prepullback.
971+
Variables (Q : cat) (A B : Q) (csp : cospan A B).
972+
(* TODO: why can't we do that? *)
973+
(* HB.instance Definition _ := Cat.on (prepullback csp). *)
974+
HB.instance Definition _ :=
975+
Quiver_IsPreCat.Build (prepullback csp)
976+
(fun a => \idmap_(a : span A B))
977+
(* TODO: study how to make this coercion trigger
978+
even when the target is not reduced to span *)
979+
(fun a b c f g => f \; g).
980+
Lemma prepullback_is_cat : PreCat_IsCat (prepullback csp).
981+
Proof. (* should be copied from the cat on span *)
982+
constructor=> [a b f|a b f|a b c d f g h];
983+
[exact: comp1o|exact: compo1|exact: compoA].
984+
Qed.
985+
End prepullback.
986+
987+
Definition pb_terminal (Q : precat)
988+
(A B : Q) (c : cospan A B) (s : prepullback c) :
989+
obj (prepullback c) := s.
990+
#[wrapper]
991+
HB.mixin Record prepullback_isTerminal (Q : precat)
992+
(A B : Q) (c : cospan A B)
993+
(s : span A B) of isPrePullback Q A B c s := {
994+
prepullback_terminal :
995+
IsTerminal (prepullback c) (pb_terminal s)
996+
}.
997+
#[short(type="pullback")]
998+
HB.structure Definition Pullback (Q : precat)
999+
(A B : Q) (c : cospan A B) :=
1000+
{s of isPrePullback Q A B c s
1001+
& prepullback_isTerminal Q A B c s }.
1002+
1003+
Notation pbsquare u v f g :=
1004+
(Pullback _ (Cospan f g) (Span u v)).
1005+
1006+
Notation "P <=> Q" := ((P -> Q) * (Q -> P))%type (at level 70).
1007+
1008+
Section th_of_pb.
1009+
Variables (Q : cat) (A B C D E F : Q).
1010+
Variables (f : A ~> D) (g : B ~> D) (h : C ~> A).
1011+
Variables (u : E ~> A) (v : E ~> B) (w : F ~> C) (z : F ~> E).
1012+
Variable (uvfg : pbsquare u v f g).
1013+
1014+
Set Printing Width 50.
1015+
Theorem pbsquarec_compP :
1016+
pbsquare w z h u <=> pbsquare w (z \; v) (h \; f) g.
1017+
Proof.
1018+
split=> [] sq.
1019+
have p : pullback (Cospan h u) := HB.pack (Span w z) sq.
1020+
1021+
Admitted.
1022+
1023+
End th_of_pb.
1024+
1025+
1026+
Module test.
1027+
1028+
Section test.
1029+
Variables (Q : precat) (A B : Q) (c : cospan A B).
1030+
Variable (p : pullback c).
1031+
Check pb_terminal p : terminal _.
1032+
1033+
1034+
1035+
1036+
1037+

0 commit comments

Comments
 (0)