partial implementation of MonoidalFunctorsCompose
[coq-categories.git] / src / MonoidalCategories_ch7_8.v
1 Generalizable All Variables.
2 Require Import Preamble.
3 Require Import Categories_ch1_3.
4 Require Import Functors_ch1_4.
5 Require Import Isomorphisms_ch1_5.
6 Require Import ProductCategories_ch1_6_1.
7 Require Import InitialTerminal_ch2_2.
8 Require Import Subcategories_ch7_1.
9 Require Import NaturalTransformations_ch7_4.
10 Require Import NaturalIsomorphisms_ch7_5.
11 Require Import Coherence_ch7_8.
12
13 (******************************************************************************)
14 (* Chapter 7.8: (Pre)Monoidal Categories                                      *)
15 (******************************************************************************)
16
17 Class BinoidalCat
18 `( C                  :  Category                               )
19 ( bin_obj'            :  C -> C -> C                            ) :=
20 { bin_obj             := bin_obj' where "a ⊗ b" := (bin_obj a b)
21 ; bin_first           :  forall a:C, Functor C C (fun x => x⊗a)
22 ; bin_second          :  forall a:C, Functor C C (fun x => a⊗x)
23 ; bin_c               := C
24 }.
25 Coercion bin_c : BinoidalCat >-> Category.
26 Notation "a ⊗ b"  := (@bin_obj _ _ _ _ _ a b)                              : category_scope.
27 Notation "C ⋊ f"  := (@fmor _ _ _ _ _ _ _ (@bin_second _ _ _ _ _ C) _ _ f) : category_scope.
28 Notation "g ⋉ C"  := (@fmor _ _ _ _ _ _ _ (@bin_first _ _ _ _ _ C) _ _ g)  : category_scope.
29 Notation "C ⋊ -"  := (@bin_second _ _ _ _ _ C)                             : category_scope.
30 Notation "- ⋉ C"  := (@bin_first _ _ _ _ _ C)                              : category_scope.
31
32 Class CentralMorphism `{BinoidalCat}`(f:a~>b) : Prop := 
33 { centralmor_first  : forall `(g:c~>d), (f ⋉ _ >>> _ ⋊ g) ~~ (_ ⋊ g >>> f ⋉ _)
34 ; centralmor_second : forall `(g:c~>d), (g ⋉ _ >>> _ ⋊ f) ~~ (_ ⋊ f >>> g ⋉ _)
35 }.
36
37 (* the central morphisms of a category constitute a subcategory *)
38 Definition Center `(bc:BinoidalCat) : SubCategory bc (fun _ => True) (fun _ _ _ _ f => CentralMorphism f).
39   apply Build_SubCategory; intros; apply Build_CentralMorphism; intros.
40   abstract (setoid_rewrite (fmor_preserves_id(bin_first c));
41               setoid_rewrite (fmor_preserves_id(bin_first d));
42               setoid_rewrite left_identity; setoid_rewrite right_identity; reflexivity).
43   abstract (setoid_rewrite (fmor_preserves_id(bin_second c));
44               setoid_rewrite (fmor_preserves_id(bin_second d));
45               setoid_rewrite left_identity; setoid_rewrite right_identity; reflexivity).
46   abstract (setoid_rewrite <- (fmor_preserves_comp(bin_first c0));
47               setoid_rewrite associativity;
48               setoid_rewrite centralmor_first;
49               setoid_rewrite <- associativity;
50               setoid_rewrite centralmor_first;
51               setoid_rewrite associativity;
52               setoid_rewrite <- (fmor_preserves_comp(bin_first d));
53               reflexivity).
54   abstract (setoid_rewrite <- (fmor_preserves_comp(bin_second d));
55               setoid_rewrite <- associativity;
56               setoid_rewrite centralmor_second;
57               setoid_rewrite associativity;
58               setoid_rewrite centralmor_second;
59               setoid_rewrite <- associativity;
60               setoid_rewrite <- (fmor_preserves_comp(bin_second c0));
61               reflexivity).
62   Qed.
63
64 Class CommutativeCat `(BinoidalCat) :=
65 { commutative_central  :  forall `(f:a~>b), CentralMorphism f
66 ; commutative_morprod  := fun `(f:a~>b)`(g:a~>b) => f ⋉ _ >>> _ ⋊ g
67 }.
68 Notation "f × g"    := (commutative_morprod f g).
69
70 Section BinoidalCat_from_Bifunctor.
71   Context `{C:Category}{Fobj}(F:Functor (C ×× C) C Fobj).
72   Definition BinoidalCat_from_Bifunctor_first (a:C) : Functor C C (fun b => Fobj (pair_obj b a)).
73   apply Build_Functor with (fmor:=(fun a0 b (f:a0~~{C}~~>b) =>
74     @fmor _ _ _ _ _ _ _ F (pair_obj a0 a) (pair_obj b a) (pair_mor (pair_obj a0 a) (pair_obj b a) f (id a)))); intros; simpl;
75     [ abstract (set (fmor_respects(F)) as q; apply q; split; simpl; auto)
76     | abstract (set (fmor_preserves_id(F)) as q; apply q)
77     | abstract (etransitivity; 
78       [ set (@fmor_preserves_comp _ _ _ _ _ _ _ F) as q; apply q
79       | set (fmor_respects(F)) as q; apply q ];
80       split; simpl; auto) ].
81   Defined.
82   Definition BinoidalCat_from_Bifunctor_second (a:C) : Functor C C (fun b => Fobj (pair_obj a b)).
83   apply Build_Functor with (fmor:=(fun a0 b (f:a0~~{C}~~>b) =>
84     @fmor _ _ _ _ _ _ _ F (pair_obj a a0) (pair_obj a b) (pair_mor (pair_obj a a0) (pair_obj a b) (id a) f))); intros;
85     [ abstract (set (fmor_respects(F)) as q; apply q; split; simpl; auto)
86     | abstract (set (fmor_preserves_id(F)) as q; apply q)
87     | abstract (etransitivity; 
88       [ set (@fmor_preserves_comp _ _ _ _ _ _ _ F) as q; apply q
89       | set (fmor_respects(F)) as q; apply q ];
90       split; simpl; auto) ].
91   Defined.
92
93   Definition BinoidalCat_from_Bifunctor : BinoidalCat C (fun a b => Fobj (pair_obj a b)).
94    refine {| bin_first  := BinoidalCat_from_Bifunctor_first
95            ; bin_second := BinoidalCat_from_Bifunctor_second
96           |}.
97    Defined.
98
99   Lemma Bifunctors_Create_Commutative_Binoidal_Categories : CommutativeCat BinoidalCat_from_Bifunctor.
100   abstract (intros; apply Build_CommutativeCat; intros; apply Build_CentralMorphism; intros; simpl; (
101     etransitivity; [ apply (fmor_preserves_comp(F)) | idtac ]; symmetry;
102     etransitivity; [ apply (fmor_preserves_comp(F)) | idtac ];
103     apply (fmor_respects(F));
104     split;
105       [ etransitivity; [ apply left_identity | symmetry; apply right_identity ]
106       | etransitivity; [ apply right_identity | symmetry; apply left_identity ] ])).
107   Defined.
108
109 End BinoidalCat_from_Bifunctor.
110
111 (* not in Awodey *)
112 Class PreMonoidalCat `(bc:BinoidalCat(C:=C))(I:C) :=
113 { pmon_I          := I
114 ; pmon_bin        := bc
115 ; pmon_cat        := C
116 ; pmon_assoc      : forall a b, (bin_second a >>>> bin_first b) <~~~> (bin_first b >>>> bin_second a)
117 ; pmon_cancelr    :                               (bin_first I) <~~~> functor_id C
118 ; pmon_cancell    :                              (bin_second I) <~~~> functor_id C
119 ; pmon_pentagon   : Pentagon (fun a b c f => f ⋉ c) (fun a b c f => c ⋊ f) (fun a b c => #((pmon_assoc a c) b))
120 ; pmon_triangle   : Triangle (fun a b c f => f ⋉ c) (fun a b c f => c ⋊ f) (fun a b c => #((pmon_assoc a c) b))
121                              (fun a => #(pmon_cancell a)) (fun a => #(pmon_cancelr a))
122 ; pmon_assoc_rr   :  forall a b, (bin_first  (a⊗b)) <~~~> (bin_first  a >>>> bin_first  b)
123 ; pmon_assoc_ll   :  forall a b, (bin_second (a⊗b)) <~~~> (bin_second b >>>> bin_second a)
124 ; pmon_coherent_r :  forall a c d:C,  #(pmon_assoc_rr c d a) ~~ #(pmon_assoc a d c)⁻¹
125 ; pmon_coherent_l :  forall a c d:C,  #(pmon_assoc_ll c a d) ~~ #(pmon_assoc c d a)
126 }.
127 (*
128  * Premonoidal categories actually have three associators (the "f"
129  * indicates the position in which the operation is natural:
130  *
131  *  assoc    : (A ⋊ f) ⋉ C <->  A ⋊ (f ⋉  C)
132  *  assoc_rr : (f ⋉ B) ⋉ C <->  f ⋉ (B  ⊗ C)
133  *  assoc_ll : (A ⋊ B) ⋊ f <-> (A ⊗  B) ⋊ f
134  *
135  * Fortunately, in a monoidal category these are all the same natural
136  * isomorphism (and in any case -- monoidal or not -- the objects in
137  * the left column are all the same and the objects in the right
138  * column are all the same).  This formalization assumes that is the
139  * case even for premonoidal categories with non-central maps, in
140  * order to keep the complexity manageable.  I don't know much about
141  * the consequences of having them and letting them be different; you
142  * might need extra versions of the triangle/pentagon diagrams.
143  *)
144
145 Implicit Arguments pmon_cancell [ Ob Hom C bin_obj' bc I ].
146 Implicit Arguments pmon_cancelr [ Ob Hom C bin_obj' bc I ].
147 Implicit Arguments pmon_assoc   [ Ob Hom C bin_obj' bc I ].
148 Coercion pmon_bin : PreMonoidalCat >-> BinoidalCat.
149
150 (* this turns out to be Exercise VII.1.1 from Mac Lane's CWM *)
151 Lemma MacLane_ex_VII_1_1 `{mn:PreMonoidalCat(I:=EI)} a b
152   : #((pmon_cancelr mn) (a ⊗ b)) ~~ #((pmon_assoc mn a EI) b) >>> (a ⋊-) \ #((pmon_cancelr mn) b).
153   set (pmon_pentagon EI EI a b) as penta. unfold pmon_pentagon in penta.
154   set (pmon_triangle a b) as tria. unfold pmon_triangle in tria.
155   apply (fmor_respects(bin_second EI)) in tria.
156   set (@fmor_preserves_comp) as fpc.
157   setoid_rewrite <- fpc in tria.
158   set (ni_commutes (pmon_assoc mn a b)) as xx.
159   (* FIXME *)
160   Admitted.
161
162 (* Formalized Definition 3.10 *)
163 Class PreMonoidalFunctor
164 `(PM1:PreMonoidalCat(C:=C1)(I:=I1))
165 `(PM2:PreMonoidalCat(C:=C2)(I:=I2))
166  (fobj : C1 -> C2          ) :=
167 { mf_F                :> Functor C1 C2 fobj
168 ; mf_preserves_i      :  mf_F I1 ≅ I2
169 ; mf_preserves_first  :  forall a,   bin_first a >>>> mf_F  <~~~>  mf_F >>>> bin_first  (mf_F a)
170 ; mf_preserves_second :  forall a,  bin_second a >>>> mf_F  <~~~>  mf_F >>>> bin_second (mf_F a)
171 ; mf_preserves_center :  forall `(f:a~>b), CentralMorphism f -> CentralMorphism (mf_F \ f)
172 }.
173 Coercion mf_F : PreMonoidalFunctor >-> Functor.
174
175 (*******************************************************************************)
176 (* Braided and Symmetric Categories                                            *)
177
178 Class BraidedCat `(mc:PreMonoidalCat) :=
179 { br_niso        : forall a, bin_first a <~~~> bin_second a
180 ; br_swap        := fun a b => ni_iso (br_niso b) a
181 ; triangleb      : forall a:C,     #(pmon_cancelr mc a) ~~ #(br_swap a (pmon_I(PreMonoidalCat:=mc))) >>> #(pmon_cancell mc a)
182 ; hexagon1       : forall {a b c}, #(pmon_assoc mc _ _ _) >>> #(br_swap a _) >>> #(pmon_assoc mc _ _ _)
183                                    ~~ #(br_swap _ _) ⋉ c >>> #(pmon_assoc mc _ _ _) >>> b ⋊ #(br_swap _ _)
184 ; hexagon2       : forall {a b c}, #(pmon_assoc mc _ _ _)⁻¹ >>> #(br_swap _ c) >>> #(pmon_assoc mc _ _ _)⁻¹
185                                    ~~ a ⋊ #(br_swap _ _) >>> #(pmon_assoc mc _ _ _)⁻¹ >>> #(br_swap _ _) ⋉ b
186 }.
187
188 Class SymmetricCat `(bc:BraidedCat) :=
189 { symcat_swap  :  forall a b:C, #((br_swap(BraidedCat:=bc)) a b) ~~ #(br_swap _ _)⁻¹
190 }.
191
192 (* Definition 7.23 *)
193 Class MonoidalCat `{C:Category}{Fobj:prod_obj C C -> C}{F:Functor (C ×× C) C Fobj}(I:C) :=
194 { mon_f         := F
195 ; mon_i         := I
196 ; mon_c         := C
197 ; mon_first     := fun a b c (f:a~>b) => F \ pair_mor (pair_obj a c) (pair_obj b c) f (id c)
198 ; mon_second    := fun a b c (f:a~>b) => F \ pair_mor (pair_obj c a) (pair_obj c b) (id c) f
199 ; mon_cancelr   :  (func_rlecnac I >>>> F) <~~~> functor_id C
200 ; mon_cancell   :  (func_llecnac I >>>> F) <~~~> functor_id C
201 ; mon_assoc     :  ((F **** (functor_id C)) >>>> F) <~~~> func_cossa >>>> ((((functor_id C) **** F) >>>> F))
202 ; mon_pentagon  :  Pentagon mon_first mon_second (fun a b c => #(mon_assoc (pair_obj (pair_obj a b) c)))
203 ; mon_triangle  :  Triangle mon_first mon_second (fun a b c => #(mon_assoc (pair_obj (pair_obj a b) c)))
204                                                  (fun a => #(mon_cancell a)) (fun a => #(mon_cancelr a))
205 }.
206 (* Coq manual on coercions: ... only the oldest one is valid and the
207  * others are ignored. So the order of declaration of coercions is
208  * important. *)
209 Coercion mon_c   : MonoidalCat >-> Category.
210 Coercion mon_f   : MonoidalCat >-> Functor.
211 Implicit Arguments mon_f [Ob Hom C Fobj F I].
212 Implicit Arguments mon_i [Ob Hom C Fobj F I].
213 Implicit Arguments mon_c [Ob Hom C Fobj F I].
214 Implicit Arguments MonoidalCat [Ob Hom ].
215
216 Section MonoidalCat_is_PreMonoidal.
217   Context `(M:MonoidalCat).
218   Definition mon_bin_M := BinoidalCat_from_Bifunctor (mon_f M).
219   Existing Instance mon_bin_M.
220   Lemma mon_pmon_assoc : forall a b, (bin_second a >>>> bin_first b) <~~~> (bin_first b >>>> bin_second a).
221     intros.
222     set (fun c => mon_assoc (pair_obj (pair_obj a c) b)) as qq.
223     simpl in qq.
224     apply Build_NaturalIsomorphism with (ni_iso:=qq).
225     abstract (intros; set ((ni_commutes mon_assoc) (pair_obj (pair_obj a A) b) (pair_obj (pair_obj a B) b)
226                     (pair_mor (pair_obj (pair_obj a A) b) (pair_obj (pair_obj a B) b)
227                     (pair_mor (pair_obj a A) (pair_obj a B) (id a) f) (id b))) as qr;
228              apply qr).
229     Defined.
230
231   Lemma mon_pmon_assoc_rr   :  forall a b, (bin_first  (a⊗b)) <~~~> (bin_first  a >>>> bin_first  b).
232     intros.
233     set (fun c:C => mon_assoc (pair_obj (pair_obj c a) b)) as qq.
234     simpl in qq.
235     apply ni_inv.
236     apply Build_NaturalIsomorphism with (ni_iso:=qq).
237     abstract (intros; set ((ni_commutes mon_assoc) (pair_obj (pair_obj _ _) _) (pair_obj (pair_obj _ _) _)
238                     (pair_mor (pair_obj (pair_obj _ _) _) (pair_obj (pair_obj _ _) _)
239                     (pair_mor (pair_obj _ _) (pair_obj _ _) f (id a)) (id b))) as qr;
240               etransitivity; [ idtac | apply qr ];
241               apply comp_respects; try reflexivity;
242               unfold mon_f;
243               simpl;
244               apply ((fmor_respects F) (pair_obj _ _) (pair_obj _ _));
245               split; try reflexivity;
246               symmetry;
247               simpl;
248               set (@fmor_preserves_id _ _ _ _ _ _ _ F (pair_obj a b)) as qqqq;
249               simpl in qqqq;
250               apply qqqq).
251    Defined.
252
253   Lemma mon_pmon_assoc_ll   :  forall a b, (bin_second (a⊗b)) <~~~> (bin_second b >>>> bin_second a).
254     intros.
255     set (fun c:C => mon_assoc (pair_obj (pair_obj a b) c)) as qq.
256     simpl in qq.
257     set (@Build_NaturalIsomorphism _ _ _ _ _ _ _ _ (Fobj (pair_obj a b) ⋊-) (b ⋊- >>>> a ⋊-)) as qqq.
258     set (qqq qq) as q'.
259     apply q'.
260     clear q'.
261     clear qqq.
262     abstract (intros; set ((ni_commutes mon_assoc) (pair_obj (pair_obj _ _) _) (pair_obj (pair_obj _ _) _)
263                     (pair_mor (pair_obj (pair_obj _ _) _) (pair_obj (pair_obj _ _) _)
264                     (pair_mor (pair_obj _ _) (pair_obj _ _) (id a) (id b)) f)) as qr;
265               etransitivity; [ apply qr | idtac ];
266               apply comp_respects; try reflexivity;
267               unfold mon_f;
268               simpl;
269               apply ((fmor_respects F) (pair_obj _ _) (pair_obj _ _));
270               split; try reflexivity;
271               simpl;
272               set (@fmor_preserves_id _ _ _ _ _ _ _ F (pair_obj a b)) as qqqq;
273               simpl in qqqq;
274               apply qqqq).
275     Defined.
276
277   Lemma mon_pmon_cancelr : (bin_first I0) <~~~> functor_id C.
278     set (@Build_NaturalIsomorphism _ _ _ _ _ _ _ _ (bin_first I0) (functor_id C)) as qq.
279     set (mon_cancelr) as z.
280     simpl in z.
281     simpl in qq.
282     set (qq z) as zz.
283     apply zz.
284     abstract (intros;
285               set (ni_commutes mon_cancelr) as q; simpl in *;
286               apply q).
287     Defined.
288
289   Lemma mon_pmon_cancell : (bin_second I0) <~~~> functor_id C.
290     set (@Build_NaturalIsomorphism _ _ _ _ _ _ _ _ (bin_second I0) (functor_id C)) as qq.
291     set (mon_cancell) as z.
292     simpl in z.
293     simpl in qq.
294     set (qq z) as zz.
295     apply zz.
296     abstract (intros;
297               set (ni_commutes mon_cancell) as q; simpl in *;
298               apply q).
299     Defined.
300
301   Lemma mon_pmon_triangle : forall a b, #(mon_pmon_cancelr a) ⋉ b ~~ #(mon_pmon_assoc _ _ _) >>> a ⋊ #(mon_pmon_cancell b).
302     intros.
303     set mon_triangle as q.
304     simpl in q.
305     apply q.
306     Qed.
307
308   Lemma mon_pmon_pentagon a b c d : (#(mon_pmon_assoc a c b ) ⋉ d)  >>>
309                                      #(mon_pmon_assoc a d _ )   >>>
310                                 (a ⋊ #(mon_pmon_assoc b d c))
311                                   ~~ #(mon_pmon_assoc _ d c )   >>>
312                                      #(mon_pmon_assoc a _ b ).
313     set (@pentagon _ _ _ _ _ _ _ mon_pentagon) as x.
314     simpl in x.
315     unfold bin_obj.
316     unfold mon_first in x.
317     simpl in *.
318     apply x.
319     Qed.
320
321   Definition MonoidalCat_is_PreMonoidal : PreMonoidalCat (BinoidalCat_from_Bifunctor (mon_f M)) (mon_i M).
322     refine {| pmon_assoc                 := mon_pmon_assoc
323             ; pmon_cancell               := mon_pmon_cancell
324             ; pmon_cancelr               := mon_pmon_cancelr
325             ; pmon_triangle              := {| triangle := mon_pmon_triangle |}
326             ; pmon_pentagon              := {| pentagon := mon_pmon_pentagon |}
327             ; pmon_assoc_ll              := mon_pmon_assoc_ll
328             ; pmon_assoc_rr              := mon_pmon_assoc_rr
329             |}.
330     abstract (set (coincide mon_triangle) as qq; simpl in *; apply qq).
331     abstract (intros; simpl; reflexivity).
332     abstract (intros; simpl; reflexivity).
333     Defined.
334
335   Lemma MonoidalCat_all_central : forall a b (f:a~~{M}~~>b), CentralMorphism f.
336     intros;
337     set (@fmor_preserves_comp _ _ _ _ _ _ _ M) as fc.
338     apply Build_CentralMorphism;
339     intros; simpl in *.
340     etransitivity.
341       apply fc.
342       symmetry.
343       etransitivity.
344       apply fc.
345       apply (fmor_respects M).
346       simpl.
347       setoid_rewrite left_identity;
348       setoid_rewrite right_identity;
349       split; reflexivity.
350     etransitivity.
351       apply fc.
352       symmetry.
353       etransitivity.
354       apply fc.
355       apply (fmor_respects M).
356       simpl.
357       setoid_rewrite left_identity;
358       setoid_rewrite right_identity;
359       split; reflexivity.
360     Qed.
361
362 End MonoidalCat_is_PreMonoidal.
363
364 Hint Extern 1 => apply MonoidalCat_all_central.
365 Coercion MonoidalCat_is_PreMonoidal : MonoidalCat >-> PreMonoidalCat.
366
367 (* Later: generalized to premonoidal categories *)
368 Class DiagonalCat `(mc:MonoidalCat(F:=F)) :=
369 {  copy_nt      :  forall a, functor_id _ ~~~> func_diagonal >>>> F
370 ;  copy         :  forall (a:mc),   a~~{mc}~~>(bin_obj(BinoidalCat:=mc) a a)
371                 := fun a => nt_component _ _ (copy_nt a) a
372 (* for non-cartesian braided diagonal categories we also need: copy >> swap == copy *)
373 }.
374
375 (* TO DO: show that the endofunctors on any given category form a monoidal category *)
376 Section MonoidalFunctor.
377   Context `(m1:MonoidalCat(C:=C1)) `(m2:MonoidalCat(C:=C2)).
378   Class MonoidalFunctor {Mobj:C1->C2} (mf_F:Functor C1 C2 Mobj) :=
379   { mf_f         := mf_F where "f ⊕⊕ g" := (@fmor _ _ _ _ _ _ _ m2 _ _ (pair_mor (pair_obj _ _) (pair_obj _ _) f g))
380   ; mf_coherence :  (mf_F **** mf_F) >>>> (mon_f m2) <~~~> (mon_f m1) >>>> mf_F
381   ; mf_phi       := fun a b => #(mf_coherence (pair_obj a b))
382   ; mf_id        :  (mon_i m2) ≅ (mf_F (mon_i m1))
383   ; mf_cancelr   :  forall a,    #(mon_cancelr(MonoidalCat:=m2) (mf_F a)) ~~
384                                   (id (mf_F a)) ⊕⊕ #mf_id >>> mf_phi a (mon_i _) >>> mf_F \ #(mon_cancelr a)
385   ; mf_cancell   :  forall b,    #(mon_cancell (mf_F b)) ~~
386                                  #mf_id ⊕⊕ (id (mf_F b)) >>> mf_phi (mon_i _) b >>> mf_F \ #(mon_cancell b)
387   ; mf_assoc     :  forall a b c, (mf_phi a b) ⊕⊕ (id (mf_F c)) >>> (mf_phi _ c) >>>
388                                   (mf_F \ #(mon_assoc (pair_obj (pair_obj a b) c) )) ~~
389                                           #(mon_assoc (pair_obj (pair_obj _ _) _) )  >>>
390                                   (id (mf_F a)) ⊕⊕ (mf_phi b c) >>> (mf_phi a _)
391   }.
392 End MonoidalFunctor.
393 Coercion mf_f : MonoidalFunctor >-> Functor.
394 Implicit Arguments mf_coherence [ Ob Hom C1 Fobj F I0 m1 Ob0 Hom0 C2 Fobj0 F0 I1 m2 Mobj mf_F ].
395 Implicit Arguments mf_id        [ Ob Hom C1 Fobj F I0 m1 Ob0 Hom0 C2 Fobj0 F0 I1 m2 Mobj mf_F ].
396
397 Section MonoidalFunctorsCompose.
398   Context `(m1:MonoidalCat).
399   Context `(m2:MonoidalCat).
400   Context `(m3:MonoidalCat).
401   Context  {f1obj}(f1:@Functor _ _ m1 _ _ m2 f1obj).
402   Context  {f2obj}(f2:@Functor _ _ m2 _ _ m3 f2obj).
403   Context  (mf1:MonoidalFunctor m1 m2 f1).
404   Context  (mf2:MonoidalFunctor m2 m3 f2).
405
406   Lemma mf_compose_coherence : (f1 >>>> f2) **** (f1 >>>> f2) >>>> m3 <~~~> m1 >>>> (f1 >>>> f2).
407     set (mf_coherence mf1) as mc1.
408     set (mf_coherence mf2) as mc2.
409     set (@ni_comp) as q.
410     set (q _ _ _ _ _ _ _ ((f1 >>>> f2) **** (f1 >>>> f2) >>>> m3) _ ((f1 **** f1 >>>> m2) >>>> f2) _ (m1 >>>> (f1 >>>> f2))) as qq.
411     apply qq; clear qq; clear q.
412     apply (@ni_comp _ _ _ _ _ _ _ _ _ (f1 **** f1 >>>> (f2 **** f2 >>>> m3)) _ _).
413     apply (@ni_comp _ _ _ _ _ _ _ _ _ ((f1 **** f1 >>>> f2 **** f2) >>>> m3) _ _).
414     eapply ni_respects.
415       apply ni_prod_comp.
416       apply ni_id.
417     apply ni_associativity.
418     apply ni_inv.
419     eapply ni_comp.
420       apply (ni_associativity (f1 **** f1) m2 f2).
421       apply (ni_respects (F0:=f1 **** f1)(F1:=f1 **** f1)(G0:=(m2 >>>> f2))(G1:=(f2 **** f2 >>>> m3))).
422         apply ni_id.
423         apply ni_inv.
424         apply mc2.
425     apply ni_inv.
426     eapply ni_comp.
427       eapply ni_inv.
428       apply (ni_associativity m1 f1 f2).
429       apply ni_respects.
430         apply ni_inv.
431         apply mc1.
432         apply ni_id.
433     Qed.
434
435   Instance MonoidalFunctorsCompose : MonoidalFunctor m1 m3 (f1 >>>> f2) :=
436   { mf_id        := id_comp         (mf_id mf2) (functors_preserve_isos f2 (mf_id mf1))
437   ; mf_coherence := mf_compose_coherence
438   }.
439   admit.
440   admit.
441   admit.
442   Defined.
443
444 End MonoidalFunctorsCompose.
445
446 Class CartesianCat `(mc:MonoidalCat) :=
447 { car_terminal  : Terminal mc
448 ; car_one       : (@one _ _ _ car_terminal) ≅ (mon_i mc)
449 ; car_diagonal  : DiagonalCat mc
450 ; car_law1      : forall {a}, id a ~~ (copy(DiagonalCat:=car_diagonal) a) >>> ((drop a >>> #car_one) ⋉ a) >>> (#(pmon_cancell mc _))
451 ; car_law2      : forall {a}, id a ~~ (copy(DiagonalCat:=car_diagonal) a) >>> (a ⋊ (drop a >>> #car_one)) >>> (#(pmon_cancelr mc _))
452 ; car_mn        := mc
453 }.
454 Coercion car_diagonal : CartesianCat >-> DiagonalCat.
455 Coercion car_terminal : CartesianCat >-> Terminal.
456 Coercion car_mn       : CartesianCat >-> MonoidalCat.
457
458 Section CenterMonoidal.
459
460   Context `(mc:PreMonoidalCat(I:=pI)).
461
462   Definition CenterMonoidal_Fobj : (Center mc) ×× (Center mc) -> Center mc.
463     intro.
464     destruct X as [a b].
465     destruct a as [a apf].
466     destruct b as [b bpf].
467     exists (a ⊗ b); auto.
468     Defined.
469
470   Definition CenterMonoidal_F : Functor _ _ CenterMonoidal_Fobj.
471     admit.
472     Defined.
473
474   Definition CenterMonoidal  : MonoidalCat _ _ CenterMonoidal_F (exist _ pI I).
475     admit.
476     Defined.
477
478 End CenterMonoidal.
479