almost finished with main theorem
[coq-hetmet.git] / src / GeneralizedArrowFromReification.v
1 (*********************************************************************************************************************************)
2 (* GeneralizedArrowFromReification:                                                                                              *)
3 (*                                                                                                                               *)
4 (*   Turn a generalized arrow into a reification                                                                                 *)
5 (*                                                                                                                               *)
6 (*********************************************************************************************************************************)
7
8 Generalizable All Variables.
9 Require Import Preamble.
10 Require Import General.
11 Require Import Categories_ch1_3.
12 Require Import Functors_ch1_4.
13 Require Import Isomorphisms_ch1_5.
14 Require Import ProductCategories_ch1_6_1.
15 Require Import OppositeCategories_ch1_6_2.
16 Require Import Enrichment_ch2_8.
17 Require Import Subcategories_ch7_1.
18 Require Import NaturalTransformations_ch7_4.
19 Require Import NaturalIsomorphisms_ch7_5.
20 Require Import MonoidalCategories_ch7_8.
21 Require Import Coherence_ch7_8.
22 Require Import Enrichment_ch2_8.
23 Require Import RepresentableStructure_ch7_2.
24 Require Import Reification.
25 Require Import GeneralizedArrow.
26
27 Section GArrowFromReification.
28
29   Context  `(K:SurjectiveEnrichment ke) `(C:MonicMonoidalEnrichment ce cme) (reification : Reification K C (me_i C)).
30
31   Fixpoint garrow_fobj_ vk : C :=
32     match vk with
33     | T_Leaf None     => me_i C
34     | T_Leaf (Some a) => match a with (a1,a2) => reification_r reification a1 a2 end
35     | t1,,t2          => me_f C (pair_obj (garrow_fobj_ t1) (garrow_fobj_ t2))
36     end.
37
38   Definition garrow_fobj vk := garrow_fobj_ (projT1 (se_decomp _ K vk)).
39
40   Definition homset_tensor_iso
41     : forall vk:enr_v_mon K, (reification_rstar reification vk) ≅ ehom(ECategory:=C) (me_i C) (garrow_fobj vk).
42     intros.
43     unfold garrow_fobj.
44     set (se_decomp _ K  vk) as sevk.
45     destruct sevk.
46     simpl in *.
47     rewrite e.
48     clear e.
49     induction x; simpl.
50
51     destruct a.
52       destruct p.
53
54       apply iso_inv.
55         apply (ni_iso (reification_commutes reification e) e0).
56
57       eapply id_comp.
58         apply iso_inv.
59         apply (mf_id (reification_rstar reification)).
60         apply (mf_id (me_mf C)).
61
62       eapply id_comp.
63         apply iso_inv.
64           apply (ni_iso (mf_coherence (reification_rstar reification)) (pair_obj _ _)).
65         eapply id_comp.
66           Focus 2.
67           apply (ni_iso (mf_coherence (me_mf C)) (pair_obj _ _)).
68           unfold bin_obj.
69           apply (functors_preserve_isos (enr_v_f C) (a:=(pair_obj _ _))(b:=(pair_obj _ _))).
70           apply (iso_prod IHx1 IHx2).
71         Defined.
72
73   Definition garrow_fobj' (vk:enr_v_mon K) : FullImage (RepresentableFunctor C (me_i C)).
74     exists (ehom(ECategory:=C) (me_i C) (garrow_fobj vk)).
75     abstract (exists (garrow_fobj vk); auto).
76     Defined.
77
78   Definition step1_mor {a b}(f:a~~{enr_v_mon K}~~>b) : (garrow_fobj' a)~~{FullImage (RepresentableFunctor C (me_i C))}~~>(garrow_fobj' b).
79     exists (iso_backward (homset_tensor_iso a) 
80         >>> reification_rstar reification \ f
81         >>> iso_forward (homset_tensor_iso  b)).
82     abstract (auto).
83     Defined.
84
85   Definition step1_functor : Functor (enr_v_mon K) (FullImage (RepresentableFunctor C (me_i C))) garrow_fobj'.
86     refine {| fmor := fun a b f => step1_mor f |}.
87     abstract (intros; unfold step1_mor; simpl;
88               apply comp_respects; try reflexivity;
89               apply comp_respects; try reflexivity;
90               apply fmor_respects; auto).
91     abstract (intros; unfold step1_mor; simpl;
92       setoid_rewrite fmor_preserves_id;
93       setoid_rewrite right_identity;
94       apply iso_comp2).
95     abstract (intros;
96       unfold step1_mor;
97       simpl;
98       repeat setoid_rewrite <- associativity;
99       apply comp_respects; try reflexivity;
100       repeat setoid_rewrite associativity;
101       apply comp_respects; try reflexivity;
102       setoid_rewrite juggle2;
103       set (iso_comp1 (homset_tensor_iso b)) as qqq;
104       setoid_rewrite qqq;
105       clear qqq;
106       setoid_rewrite right_identity;
107       apply (fmor_preserves_comp reification)).
108       Defined.
109
110   Definition step1_niso : reification ≃ step1_functor >>>> InclusionFunctor _ (FullImage (RepresentableFunctor C (me_i C))).
111     exists (fun c1 => homset_tensor_iso c1).
112     abstract (intros;
113               simpl;
114               repeat setoid_rewrite <- associativity;
115               setoid_rewrite iso_comp1;
116               setoid_rewrite left_identity;
117               reflexivity).
118     Qed.
119     Opaque homset_tensor_iso.
120
121   Definition step2_functor := ff_functor_section_functor _ (ffme_mf_full C) (ffme_mf_faithful C).
122
123   Definition garrow_functor := step1_functor >>>> step2_functor.
124
125   Lemma garrow_functor_monoidal_niso
126     : (garrow_functor **** garrow_functor) >>>> (mon_f C) <~~~> (mon_f (enr_v_mon K)) >>>> garrow_functor.
127     admit.
128     Defined.
129   Lemma garrow_functor_monoidal_iso
130     : mon_i C ≅ garrow_functor (mon_i (enr_v_mon K)).
131     admit.
132     Defined.
133
134   Instance garrow_functor_monoidal : MonoidalFunctor (enr_v_mon K) C garrow_functor :=
135     { mf_coherence := garrow_functor_monoidal_niso
136     ; mf_id        := garrow_functor_monoidal_iso
137     }.
138     admit.
139     admit.
140     admit.
141     Defined.
142
143   Definition garrow_from_reification : GeneralizedArrow K C.
144     refine
145       {| ga_functor          := garrow_functor
146        ; ga_functor_monoidal := garrow_functor_monoidal
147       |}.
148     Defined.
149
150 End GArrowFromReification.
151 Opaque homset_tensor_iso.