52a3d2ab9502d135ad59f8f4a3dfa74e24aceaf3
[coq-categories.git] / src / Enrichment_ch2_8.v
1 Generalizable All Variables.
2 Require Import Preamble.
3 Require Import General.
4 Require Import Categories_ch1_3.
5 Require Import Functors_ch1_4.
6 Require Import Isomorphisms_ch1_5.
7 Require Import ProductCategories_ch1_6_1.
8 Require Import InitialTerminal_ch2_2.
9 Require Import Subcategories_ch7_1.
10 Require Import NaturalTransformations_ch7_4.
11 Require Import NaturalIsomorphisms_ch7_5.
12 Require Import Coherence_ch7_8.
13 Require Import BinoidalCategories.
14 Require Import PreMonoidalCategories.
15 Require Import MonoidalCategories_ch7_8.
16
17 (******************************************************************************)
18 (* Chapter 2.8: Hom Sets and Enriched Categories                              *)
19 (******************************************************************************)
20
21 Class ECategory `(mn:PreMonoidalCat(bc:=bc)(C:=V)(I:=EI))(Eob:Type)(Ehom:Eob->Eob->V) :=
22 { ehom               :=  Ehom where "a ~~> b" := (ehom a b)
23 ; eob_eob            :=  Eob
24 ; e_v                :=  mn
25 ; eid                :   forall a,          EI~>(a~~>a)
26 ; eid_central        :   forall a,          CentralMorphism (eid a)
27 ; ecomp              :   forall {a b c},    (a~~>b)⊗(b~~>c) ~> (a~~>c)
28 ; ecomp_central      :>  forall {a b c},    CentralMorphism (@ecomp a b c)
29 ; eleft_identity     :   forall {a b  },    eid a ⋉ (a~~>b) >>> ecomp ~~ #(pmon_cancell _)
30 ; eright_identity    :   forall {a b  },    (a~~>b) ⋊ eid b >>> ecomp ~~ #(pmon_cancelr _)
31 ; eassociativity     :   forall {a b c d},  #(pmon_assoc _ _ (_~~>_))⁻¹ >>> ecomp ⋉ (c~~>d) >>> ecomp ~~ (a~~>b) ⋊ ecomp >>> ecomp
32 }.
33 Notation "a ~~> b" := (@ehom _ _ _ _ _ _ _ _ _ _ a b) : category_scope.
34 Coercion eob_eob : ECategory >-> Sortclass.
35
36 Lemma ecomp_is_functorial `{ec:ECategory}{a b c}{x}(f:EI~~{V}~~>(a~~>b))(g:EI~~{V}~~>(b~~>c)) :
37    ((x ~~> a) ⋊-) \ (iso_backward (pmon_cancelr EI) >>> ((- ⋉EI) \ f >>> (((a ~~> b) ⋊-) \ g >>> ecomp))) >>> ecomp ~~
38    ((x ~~> a) ⋊-) \ f >>> (ecomp >>> (#(pmon_cancelr (x ~~> b)) ⁻¹ >>> (((x ~~> b) ⋊-) \ g >>> ecomp))).
39
40    set (@fmor_preserves_comp) as fmor_preserves_comp'.
41
42    (* knock off the leading (x ~~> a) ⋊ f *)
43    repeat setoid_rewrite <- associativity.
44    set (ni_commutes (@pmon_cancelr _ _ _ _ _ _ mn) f) as qq.
45    apply iso_shift_right' in qq.
46    setoid_rewrite <- associativity in qq.
47    apply symmetry in qq.
48    apply iso_shift_left' in qq.
49    apply symmetry in qq.
50    simpl in qq.
51    setoid_rewrite <- qq.
52    clear qq.
53    repeat setoid_rewrite associativity.
54    repeat setoid_rewrite <- fmor_preserves_comp'.
55    repeat setoid_rewrite associativity.
56    apply comp_respects; try reflexivity.
57
58    (* rewrite using the lemma *)
59    assert (forall {a b c x}(g:EI~~{V}~~>(b ~~> c)),
60     ((bin_second(BinoidalCat:=bc) (x ~~> a)) \ ((bin_second(BinoidalCat:=bc) (a ~~> b)) \ g))
61     ~~
62     (#(pmon_assoc (x ~~> a) _ _)⁻¹ >>>
63      (bin_second(BinoidalCat:=bc) ((x ~~> a) ⊗ (a ~~> b))) \ g >>> #(pmon_assoc (x ~~> a) _ _))).
64     symmetry.
65
66     setoid_rewrite associativity.
67     symmetry.
68     apply iso_shift_right'.
69     setoid_rewrite <- pmon_coherent_l.
70     set (ni_commutes (pmon_assoc_ll (x0~~>a0) (a0~~>b0))) as qq.
71     simpl in *.
72     apply (qq _ _ g0).
73    setoid_rewrite H.
74    clear H.
75
76    (* rewrite using eassociativity *)
77    repeat setoid_rewrite associativity.
78    set (@eassociativity _ _ _ _ _ _ _ _ _ ec x a) as qq.
79    setoid_rewrite <- qq.
80    clear qq.
81    unfold e_v.
82
83    (* knock off the trailing ecomp *)
84    repeat setoid_rewrite <- associativity.
85    apply comp_respects; try reflexivity.
86
87    (* cancel out the adjacent assoc/cossa pair *)
88    repeat setoid_rewrite associativity.
89    setoid_rewrite juggle2.
90    etransitivity.
91    apply comp_respects; [ idtac | 
92    repeat setoid_rewrite <- associativity;
93    etransitivity; [ idtac | apply left_identity ];
94    apply comp_respects; [ idtac | reflexivity ];
95    apply iso_comp1 ].
96    apply reflexivity.
97
98    (* now swap the order of ecomp⋉(b ~~> c) and ((x ~~> a) ⊗ (a ~~> b))⋊g *)
99    repeat setoid_rewrite associativity.
100    set (@centralmor_first) as se.
101    setoid_rewrite <- se.
102    clear se.
103
104    (* and knock the trailing (x ~~> b)⋊ g off *)
105    repeat setoid_rewrite <- associativity.
106    apply comp_respects; try reflexivity.
107
108    (* push the ecomp forward past the rlecnac *)
109    set (ni_commutes (@pmon_cancelr _ _ _ _ _ _ mn) (@ecomp _ _ _ _ _ _ _ _ _ ec x a b)) as qq.
110    symmetry in qq.
111    apply iso_shift_left' in qq.
112    setoid_rewrite associativity in qq.
113    symmetry in qq.
114    apply iso_shift_right' in qq.
115    simpl in qq.
116    setoid_rewrite qq.
117    clear qq.
118
119    (* and knock off the trailing ecomp *)
120    apply comp_respects; try reflexivity.
121
122    setoid_replace (iso_backward ((pmon_cancelr) ((x ~~> a) ⊗ (a ~~> b)))) with
123      (iso_backward ((pmon_cancelr) ((x ~~> a) ⊗ (a ~~> b))) >>> id _).
124    simpl.
125    set (@iso_shift_right') as ibs.
126    simpl in ibs.
127    apply ibs.
128    clear ibs.
129
130    setoid_rewrite (MacLane_ex_VII_1_1 (x~~>a) (a~~>b)).
131    setoid_rewrite juggle3.
132    set (fmor_preserves_comp ((x ~~> a) ⋊-)) as q.
133    simpl in q.
134    setoid_rewrite q.
135    clear q.
136    setoid_rewrite iso_comp1.
137    setoid_rewrite fmor_preserves_id.
138    setoid_rewrite right_identity.
139    apply iso_comp1.
140
141    (* leftovers *)
142    symmetry.
143    apply right_identity.
144    apply ecomp_central.
145    Qed.
146
147 Lemma underlying_associativity `{ec:ECategory(mn:=mn)(EI:=EI)(Eob:=Eob)(Ehom:=Ehom)} :
148    forall {a b : Eob} (f : EI ~~{ V }~~> a ~~> b) {c : Eob}
149    (g : EI ~~{ V }~~> b ~~> c) {d : Eob} (h : EI ~~{ V }~~> c ~~> d),
150    ((((#(pmon_cancelr EI) ⁻¹ >>> (f ⋉ EI >>> (a ~~> b) ⋊ g)) >>> ecomp) ⋉ EI >>> (a ~~> c) ⋊ h)) >>> ecomp ~~
151    ((f ⋉ EI >>> (a ~~> b) ⋊ ((#(pmon_cancelr EI) ⁻¹ >>> (g ⋉ EI >>> (b ~~> c) ⋊ h)) >>> ecomp))) >>> ecomp.
152
153   intros; symmetry; etransitivity;
154      [ setoid_rewrite associativity; apply comp_respects;
155         [ apply reflexivity | repeat setoid_rewrite associativity; apply (ecomp_is_functorial(x:=a) g h) ] | idtac ].
156
157   repeat setoid_rewrite <- fmor_preserves_comp.
158     repeat setoid_rewrite <- associativity.
159     apply comp_respects; try reflexivity.
160     apply comp_respects; try reflexivity.
161
162   set (ni_commutes (@pmon_cancelr _ _ _ _ _ _ mn) f) as qq.
163     apply iso_shift_right' in qq.
164     symmetry in qq.
165     setoid_rewrite <- associativity in qq.
166     apply iso_shift_left' in qq.
167     apply (fmor_respects (bin_first EI)) in qq.
168     setoid_rewrite <- fmor_preserves_comp in qq.
169     setoid_rewrite qq.
170     clear qq.
171
172   repeat setoid_rewrite <- fmor_preserves_comp.
173     repeat setoid_rewrite associativity.
174     apply comp_respects; try reflexivity.
175
176   repeat setoid_rewrite associativity.
177   set (ni_commutes (@pmon_cancelr _ _ _ _ _ _ mn) (@ecomp _ _ _ _ _ _ _ _ _ ec a b c)) as qq.
178     apply iso_shift_right' in qq.
179     symmetry in qq.
180     setoid_rewrite <- associativity in qq.
181     apply iso_shift_left' in qq.
182     symmetry in qq.
183     simpl in qq.
184     setoid_rewrite qq.
185     clear qq.
186
187   repeat setoid_rewrite <- associativity.
188     apply comp_respects; try reflexivity.
189
190   idtac.
191     set (iso_shift_left'
192          (iso_backward (pmon_cancelr (a ~~> b)) ⋉ EI >>> ((a ~~> b) ⋊ g) ⋉ EI) ((a ~~> b) ⋊ g)
193          ((pmon_cancelr ((a ~~> b) ⊗ (b ~~> c))))) as xx.
194     symmetry.
195     etransitivity; [ apply xx | apply comp_respects; try reflexivity ].
196     clear xx.
197
198   setoid_rewrite (fmor_preserves_comp (bin_first EI)).
199     set (ni_commutes (@pmon_cancelr _ _ _ _ _ _ mn) ((iso_backward (pmon_cancelr (a ~~> b)) >>> (a ~~> b) ⋊ g))) as qq.
200     simpl in qq.
201     setoid_rewrite <- qq.
202     clear qq.
203
204   setoid_rewrite <- associativity.
205     setoid_rewrite iso_comp1.
206     apply left_identity.
207
208   Qed.
209
210 Instance Underlying `(ec:ECategory(mn:=mn)(EI:=EI)(Eob:=Eob)(Ehom:=Ehom)) : Category Eob (fun a b => EI~>(a~~>b)) :=
211 { id    := fun a                                  => eid a
212 ; comp  := fun a b c g f                          => #(pmon_cancelr _)⁻¹ >>> (g ⋉ _ >>> _ ⋊ f) >>> ecomp
213 ; eqv   := fun a b (f:EI~>(a~~>b))(g:EI~>(a~~>b)) => f ~~ g
214 }.
215   abstract (intros; apply Build_Equivalence;
216     [ unfold Reflexive
217     | unfold Symmetric
218     | unfold Transitive]; intros; simpl; auto).
219   abstract (intros; unfold Proper; unfold respectful; intros; simpl;
220   repeat apply comp_respects; try apply reflexivity;
221     [ apply (fmor_respects (bin_first EI)) | idtac ]; auto;
222     apply (fmor_respects (bin_second (a~~>b))); auto).
223   abstract (intros;
224     set (fun c d => centralmor_first(c:=c)(d:=d)(CentralMorphism:=(eid_central a))) as q;
225     setoid_rewrite q;
226     repeat setoid_rewrite associativity;
227     setoid_rewrite eleft_identity;
228     setoid_rewrite <- (ni_commutes (@pmon_cancell _ _ _ _ _ _ mn));
229     setoid_rewrite <- associativity;
230     set (coincide pmon_triangle) as qq;
231     setoid_rewrite qq;
232     simpl;
233     setoid_rewrite iso_comp2;
234     apply left_identity).
235   abstract (intros;
236     repeat setoid_rewrite associativity;
237     setoid_rewrite eright_identity;
238     setoid_rewrite <- (ni_commutes (@pmon_cancelr _ _ _ _ _ _ mn));
239     setoid_rewrite <- associativity;
240     simpl;
241     setoid_rewrite iso_comp2;
242     apply left_identity).
243   abstract (intros;
244     repeat setoid_rewrite associativity;
245     apply comp_respects; try reflexivity;
246     repeat setoid_rewrite <- associativity;
247     apply underlying_associativity).
248   Defined.
249 Coercion Underlying : ECategory >-> Category.
250
251 Class EFunctor
252   `{mn:PreMonoidalCat(bc:=bc)(C:=V)(I:=EI)}
253    {Eob1}{EHom1}(ec1:ECategory mn Eob1 EHom1)
254    {Eob2}{EHom2}(ec2:ECategory mn Eob2 EHom2)
255    (EFobj : Eob1 -> Eob2)
256 :=
257 { efunc_fobj           := EFobj
258 ; efunc                :  forall a b:Eob1, (a ~~> b) ~~{V}~~> ( (EFobj a) ~~> (EFobj b) )
259 ; efunc_central        :  forall a b,      CentralMorphism (efunc a b)
260 ; efunc_preserves_id   :  forall a,        eid a >>> efunc a a ~~ eid (EFobj a)
261 ; efunc_preserves_comp :  forall a b c,    (efunc a b) ⋉ _ >>>  _ ⋊ (efunc b c) >>> ecomp ~~ ecomp >>> efunc a c
262 }.
263 Coercion efunc_fobj : EFunctor >-> Funclass.
264 Implicit Arguments efunc [ Ob Hom V bin_obj' bc EI mn Eob1 EHom1 ec1 Eob2 EHom2 ec2 EFobj ].
265
266 Definition efunctor_id 
267   `{mn:PreMonoidalCat(bc:=bc)(C:=V)(I:=EI)}
268    {Eob1}{EHom1}(ec1:ECategory mn Eob1 EHom1)
269    : EFunctor ec1 ec1 (fun x => x).
270   refine {| efunc := fun a b => id (a ~~> b) |}.
271   abstract (intros; apply Build_CentralMorphism; intros;
272     setoid_rewrite fmor_preserves_id;
273     setoid_rewrite right_identity;
274     setoid_rewrite left_identity;
275     reflexivity).
276   abstract (intros; apply right_identity).
277   abstract (intros;
278     setoid_rewrite fmor_preserves_id;
279     setoid_rewrite right_identity;
280     setoid_rewrite left_identity;
281     reflexivity).
282   Defined.
283
284 Definition efunctor_comp
285   `{mn:PreMonoidalCat(bc:=bc)(C:=V)(I:=EI)}
286    {Eob1}{EHom1}(ec1:ECategory mn Eob1 EHom1)
287    {Eob2}{EHom2}(ec2:ECategory mn Eob2 EHom2)
288    {Eob3}{EHom3}(ec3:ECategory mn Eob3 EHom3)
289    {Fobj}(F:EFunctor ec1 ec2 Fobj)
290    {Gobj}(G:EFunctor ec2 ec3 Gobj)
291    : EFunctor ec1 ec3 (Gobj ○ Fobj).
292   refine {| efunc := fun a b => (efunc F a b) >>> (efunc G _ _) |}; intros; simpl.
293   abstract (apply Build_CentralMorphism; intros;
294     [ set (fun a b c d => centralmor_first(CentralMorphism:=(efunc_central(EFunctor:=F)) a b)(c:=c)(d:=d)) as fc1
295     ; set (fun a b c d => centralmor_first(CentralMorphism:=(efunc_central(EFunctor:=G)) a b)(c:=c)(d:=d)) as gc1
296     ; setoid_rewrite <- (fmor_preserves_comp (-⋉d))
297     ; setoid_rewrite <- (fmor_preserves_comp (-⋉c))
298     ; setoid_rewrite <- associativity
299     ; setoid_rewrite <- fc1
300     ; setoid_rewrite    associativity
301     ; setoid_rewrite <- gc1
302     ; reflexivity
303     | set (fun a b c d => centralmor_second(CentralMorphism:=(efunc_central(EFunctor:=F)) a b)(c:=c)(d:=d)) as fc2
304     ; set (fun a b c d => centralmor_second(CentralMorphism:=(efunc_central(EFunctor:=G)) a b)(c:=c)(d:=d)) as gc2
305     ; setoid_rewrite <- (fmor_preserves_comp (d⋊-))
306     ; setoid_rewrite <- (fmor_preserves_comp (c⋊-))
307     ; setoid_rewrite <- associativity
308     ; setoid_rewrite    fc2
309     ; setoid_rewrite    associativity
310     ; setoid_rewrite    gc2
311     ; reflexivity ]).
312   abstract (setoid_rewrite <- associativity;
313     setoid_rewrite efunc_preserves_id;
314     setoid_rewrite efunc_preserves_id;
315     reflexivity).
316   abstract (repeat setoid_rewrite associativity;
317     set (fmor_preserves_comp (-⋉(b~~>c))) as q; setoid_rewrite <- q; clear q;
318     repeat setoid_rewrite associativity;
319     set (fmor_preserves_comp (((Gobj (Fobj a) ~~> Gobj (Fobj b))⋊-))) as q; setoid_rewrite <- q; clear q;
320     set (fun d e => centralmor_second(c:=d)(d:=e)(CentralMorphism:=(efunc_central(EFunctor:=F) b c))) as qq;
321     setoid_rewrite juggle2;
322     setoid_rewrite juggle2;
323     setoid_rewrite qq;
324     clear qq;
325     repeat setoid_rewrite associativity;
326     set ((efunc_preserves_comp(EFunctor:=G)) (Fobj a) (Fobj b) (Fobj c)) as q;
327     repeat setoid_rewrite associativity;
328     repeat setoid_rewrite associativity in q;
329     setoid_rewrite q;
330     clear q;
331     repeat setoid_rewrite <- associativity;
332     apply comp_respects; try reflexivity;
333     set ((efunc_preserves_comp(EFunctor:=F)) a b c) as q;
334     apply q).
335   Defined.
336
337 Instance UnderlyingFunctor `(EF:@EFunctor Ob Hom V bin_obj' bc EI mn Eob1 EHom1 ec1 Eob2 EHom2 ec2 Eobj)
338   : Functor (Underlying ec1) (Underlying ec2) Eobj :=
339   { fmor := fun a b (f:EI~~{V}~~>(a~~>b)) => f >>> (efunc _ a b) }.
340   abstract (intros; simpl; apply comp_respects; try reflexivity; auto).
341   abstract (intros; simpl; apply efunc_preserves_id).
342   abstract (intros;
343             simpl;
344             repeat setoid_rewrite associativity;
345             setoid_rewrite <- efunc_preserves_comp;
346             repeat setoid_rewrite associativity;
347               apply comp_respects; try reflexivity;
348             set (@fmor_preserves_comp _ _ _ _ _ _ _ (bin_first EI)) as qq;
349               setoid_rewrite <- qq;
350               clear qq;
351             repeat setoid_rewrite associativity;
352               apply comp_respects; try reflexivity;
353             repeat setoid_rewrite <- associativity;
354               apply comp_respects; try reflexivity;
355             set (@fmor_preserves_comp _ _ _ _ _ _ _ (bin_second (Eobj a ~~> Eobj b))) as qq;
356               setoid_rewrite <- qq;
357               repeat setoid_rewrite <- associativity;
358               apply comp_respects; try reflexivity;
359               clear qq;
360             apply (centralmor_first(CentralMorphism:=(efunc_central a b)))).
361   Defined.
362   Coercion UnderlyingFunctor : EFunctor >-> Functor.
363
364 Class EBinoidalCat `(ec:ECategory) :=
365 { ebc_bobj   :  ec -> ec -> ec
366 ; ebc_first  :  forall a:ec, EFunctor ec ec (fun x => ebc_bobj x a)
367 ; ebc_second :  forall a:ec, EFunctor ec ec (fun x => ebc_bobj a x)
368 ; ebc_ec     := ec  (* this isn't a coercion - avoids duplicate paths *)
369 }.
370
371 Instance EBinoidalCat_is_binoidal `(ebc:EBinoidalCat(ec:=ec)) : BinoidalCat (Underlying ec) ebc_bobj.
372   apply Build_BinoidalCat.
373   apply (fun a => UnderlyingFunctor (ebc_first a)).
374   apply (fun a => UnderlyingFunctor (ebc_second a)).
375   Defined.
376
377 Coercion EBinoidalCat_is_binoidal : EBinoidalCat >-> BinoidalCat.
378
379 (* Enriched Fibrations *)
380 Section EFibration.
381
382   Context `{E:ECategory}.
383   Context  {Eob2}{Ehom2}{B:@ECategory Ob Hom V bin_obj' mn EI mn Eob2 Ehom2}.
384   Context  {efobj}(F:EFunctor E B efobj).
385
386   (*
387    * A morphism is prone if its image factors through the image of
388    * another morphism with the same codomain just in case the factor
389    * is the image of a unique morphism.  One might say that it
390    * "uniquely reflects factoring through morphisms with the same
391    * codomain".
392    *)
393   Definition Prone {e e'}(φ:EI~~{V}~~>(e'~~>e)) :=
394   forall e'' (ψ:EI~~{V}~~>(e''~~>e)) (g:(efobj e'')~~{B}~~>(efobj e')),
395        (comp(Category:=B) _ _ _ g (φ >>> (efunc F _ _))) ~~
396        ψ >>> (efunc F _ _)
397    ->  { χ:e''~~{E}~~>e' & ψ ~~ χ >>> φ & g ~~ comp(Category:=V) _ _ _ χ (efunc F e'' e') }.
398        (* FIXME: χ must also be unique *)
399
400   (* a functor is a Street Fibration if morphisms with codomain in its image are, up to iso, the images of prone morphisms *)
401
402   (* Street was the first to define non-evil fibrations using isomorphisms (for cleavage_pf below) rather than equality *)
403   Structure StreetCleavage (e:E)(b:B)(f:b~~{B}~~>(efobj e)) :=
404   { cleavage_e'   : E
405   ; cleavage_pf   : (efobj cleavage_e') ≅ b
406   ; cleavage_phi  : cleavage_e' ~~{E}~~> e
407   ; cleavage_cart : Prone cleavage_phi
408   ; cleavage_eqv  : #cleavage_pf >>> f  ~~ comp(Category:=V) _ _ _ cleavage_phi (efunc F _ _)
409   }.
410
411   (* if V, the category enriching E and B is V-enriched, we get a functor Bop->Vint *)
412
413   (* Every equivalence of categories is a Street fibration *)
414
415   (* this is actually a "Street Fibration", the non-evil version of a Grothendieck Fibration *)
416   Definition EFibration := forall e b f, exists cl:StreetCleavage e b f, True.
417
418   Definition ClovenEFibration := forall e b f, StreetCleavage e b f.
419
420   (*
421    * Now, a language has polymorphic types iff its category of
422    * judgments contains a second enriched category, the category of
423    * Kinds, and the category of types is fibered over the category of
424    * Kinds, and the weakening functor-of-fibers has a right adjoint.
425    *
426    * http://ncatlab.org/nlab/show/Grothendieck+fibration
427    *
428    * I suppose we'll need to also ask that the R-functors takes
429    * Prone morphisms to Prone morphisms.
430    *)
431 End EFibration.
432