4a9e7908ce49a83d2626bf3a7daba72cf9d0a271
[coq-hetmet.git] / src / GeneralizedArrowFromReification.v
1 (*********************************************************************************************************************************)
2 (* GeneralizedArrowFromReification:                                                                                              *)
3 (*                                                                                                                               *)
4 (*   Turn a reification into a generalized arrow                                                                                 *)
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 (HomFunctor 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 (HomFunctor 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   (* The poorly-named "step1_functor" is a functor from the full subcategory in the range of the reification functor
86    * to the full subcategory in the range of the [host language's] Hom(I,-) functor *)
87   Definition step1_functor : Functor (enr_v_mon K) (FullImage (HomFunctor C (me_i C))) garrow_fobj'.
88     refine {| fmor := fun a b f => step1_mor f |}.
89     abstract (intros; unfold step1_mor; simpl;
90               apply comp_respects; try reflexivity;
91               apply comp_respects; try reflexivity;
92               apply fmor_respects; auto).
93     abstract (intros; unfold step1_mor; simpl;
94       setoid_rewrite fmor_preserves_id;
95       setoid_rewrite right_identity;
96       apply iso_comp2).
97     abstract (intros;
98       unfold step1_mor;
99       simpl;
100       repeat setoid_rewrite <- associativity;
101       apply comp_respects; try reflexivity;
102       repeat setoid_rewrite associativity;
103       apply comp_respects; try reflexivity;
104       setoid_rewrite juggle2;
105       set (iso_comp1 (homset_tensor_iso b)) as qqq;
106       setoid_rewrite qqq;
107       clear qqq;
108       setoid_rewrite right_identity;
109       apply (fmor_preserves_comp reification)).
110       Defined.
111
112   Definition step1_niso : reification ≃ step1_functor >>>> InclusionFunctor _ (FullImage (HomFunctor C (me_i C))).
113     exists (fun c1 => homset_tensor_iso c1).
114     abstract (intros;
115               simpl;
116               repeat setoid_rewrite <- associativity;
117               setoid_rewrite iso_comp1;
118               setoid_rewrite left_identity;
119               reflexivity).
120     Qed.
121
122   (* the "step2_functor" is the section of the Hom(I,-) functor *)
123   Definition step2_functor := ff_functor_section_functor _ (ffme_mf_full C) (ffme_mf_faithful C).
124
125   (* the generalized arrow is the composition of the two steps *)
126   Definition garrow_functor := step1_functor >>>> step2_functor.
127
128   Lemma garrow_functor_monoidal_iso_i
129     : mon_i C ≅ garrow_functor (mon_i (enr_v_mon K)).
130     admit.
131     Defined.
132
133   Lemma garrow_functor_monoidal_iso :
134     forall X Y:enr_v_mon K, 
135       garrow_functor (bin_obj(BinoidalCat:=enr_v_mon K) X Y) ≅ bin_obj(BinoidalCat:=me_mon C) (garrow_functor X) (garrow_functor Y).
136     admit.
137     Defined.
138
139   Definition garrow_functor_monoidal_niso
140     : (garrow_functor **** garrow_functor) >>>> (mon_f C) <~~~> (mon_f (enr_v_mon K)) >>>> garrow_functor.
141     admit.
142     Defined.
143     Opaque homset_tensor_iso.
144
145   Instance garrow_functor_monoidal : MonoidalFunctor (enr_v_mon K) C garrow_functor :=
146     { mf_coherence   := garrow_functor_monoidal_niso
147     ; mf_id          := garrow_functor_monoidal_iso_i
148     }.
149     admit.
150     admit.
151     admit.
152     Defined.
153
154   Definition garrow_from_reification : GeneralizedArrow K C.
155     refine
156       {| ga_functor          := garrow_functor
157        ; ga_functor_monoidal := garrow_functor_monoidal
158       |}.
159     Defined.
160
161 End GArrowFromReification.
162 Opaque homset_tensor_iso.