add commented-out definitions for analytic proofs and cut elimination
[coq-hetmet.git] / src / ProgrammingLanguage.v
1 (*********************************************************************************************************************************)
2 (* ProgrammingLanguage                                                                                                           *)
3 (*                                                                                                                               *)
4 (*   Basic assumptions about programming languages.                                                                              *)
5 (*                                                                                                                               *)
6 (*********************************************************************************************************************************)
7
8 Generalizable All Variables.
9 Require Import Preamble.
10 Require Import General.
11 Require Import Categories_ch1_3.
12 Require Import InitialTerminal_ch2_2.
13 Require Import Functors_ch1_4.
14 Require Import Isomorphisms_ch1_5.
15 Require Import ProductCategories_ch1_6_1.
16 Require Import OppositeCategories_ch1_6_2.
17 Require Import Enrichment_ch2_8.
18 Require Import Subcategories_ch7_1.
19 Require Import NaturalTransformations_ch7_4.
20 Require Import NaturalIsomorphisms_ch7_5.
21 Require Import BinoidalCategories.
22 Require Import PreMonoidalCategories.
23 Require Import MonoidalCategories_ch7_8.
24 Require Import Coherence_ch7_8.
25 Require Import Enrichment_ch2_8.
26 Require Import RepresentableStructure_ch7_2.
27 Require Import FunctorCategories_ch7_7.
28
29 Require Import Enrichments.
30 Require Import NaturalDeduction.
31 Require Import NaturalDeductionCategory.
32
33 Section Programming_Language.
34
35   Context {T    : Type}.               (* types of the language *)
36
37   Definition PLJudg := (Tree ??T) * (Tree ??T).
38   Definition sequent := @pair (Tree ??T) (Tree ??T).
39      Notation "cs |= ss" := (sequent cs ss) : pl_scope.
40
41   Context {Rule : Tree ??PLJudg -> Tree ??PLJudg -> Type}.
42
43   Notation "H /⋯⋯/ C" := (ND Rule H C) : pl_scope.
44
45   Open Scope pf_scope.
46   Open Scope nd_scope.
47   Open Scope pl_scope.
48
49   Class ProgrammingLanguage :=
50   { pl_eqv0               :  @ND_Relation PLJudg Rule
51   ; pl_snd                :> @SequentND PLJudg Rule _ sequent
52   ; pl_cnd                :> @ContextND PLJudg Rule T sequent pl_snd
53   ; pl_eqv1               :> @SequentND_Relation PLJudg Rule _ sequent pl_snd pl_eqv0
54   ; pl_eqv                :> @ContextND_Relation PLJudg Rule _ sequent pl_snd pl_cnd pl_eqv0 pl_eqv1
55   }.
56   Notation "pf1 === pf2" := (@ndr_eqv _ _ pl_eqv _ _ pf1 pf2) : temporary_scope3.
57
58   Section LanguageCategory.
59
60     Context (PL:ProgrammingLanguage).
61
62     (* category of judgments in a fixed type/coercion context *)
63     Definition Judgments_cartesian := @Judgments_Category_CartesianCat _ Rule pl_eqv.
64
65     Definition JudgmentsL          := Judgments_cartesian.
66
67     Definition identityProof t : [] ~~{JudgmentsL}~~> [t |= t].
68       unfold hom; simpl.
69       apply snd_initial.
70       Defined.
71
72     Definition cutProof a b c : [a |= b],,[b |= c] ~~{JudgmentsL}~~> [a |= c].
73       unfold hom; simpl.
74       apply snd_cut.
75       Defined.
76
77     Existing Instance pl_eqv.
78
79     Definition TypesL : ECategory JudgmentsL (Tree ??T) (fun x y => [x|=y]).
80       refine
81       {| eid   := identityProof
82        ; ecomp := cutProof
83       |}; intros.
84       apply (mon_commutative(MonoidalCat:=JudgmentsL)).
85       apply (mon_commutative(MonoidalCat:=JudgmentsL)).
86       unfold identityProof; unfold cutProof; simpl; eapply cndr_inert. apply pl_eqv. auto. auto.
87       unfold identityProof; unfold cutProof; simpl; eapply cndr_inert. apply pl_eqv. auto. auto.
88       unfold identityProof; unfold cutProof; simpl; eapply cndr_inert. apply pl_eqv. auto. auto.
89       apply ndpc_comp; auto.
90       apply ndpc_comp; auto.
91       Defined.
92
93     Instance Types_first c : EFunctor TypesL TypesL (fun x => x,,c ) :=
94       { efunc := fun x y => cnd_expand_right(ContextND:=pl_cnd) x y c }.
95       intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)).
96       intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof.
97       apply (cndr_inert pl_cnd); auto.
98       intros. unfold ehom. unfold comp. simpl. unfold cutProof.
99       rewrite <- (@ndr_prod_preserves_comp _ _ pl_eqv _ _ (cnd_expand_right _ _ c) _ _ (nd_id1 (b|=c0))
100                   _ (nd_id1 (a,,c |= b,,c))  _ (cnd_expand_right _ _ c)).
101       setoid_rewrite (@ndr_comp_right_identity _ _ pl_eqv _ [a,, c |= b,, c]).
102       setoid_rewrite (@ndr_comp_left_identity  _ _ pl_eqv [b |= c0]).
103       simpl; eapply cndr_inert. apply pl_eqv. auto. auto.
104       Defined.
105
106     Instance Types_second c : EFunctor TypesL TypesL (fun x => c,,x) :=
107       { efunc := fun x y => ((@cnd_expand_left _ _ _ _ _ _ x y c)) }.
108       intros; apply (mon_commutative(MonoidalCat:=JudgmentsL)).
109       intros. unfold ehom. unfold hom. unfold identityProof. unfold eid. simpl. unfold identityProof.
110       eapply cndr_inert; auto. apply pl_eqv.
111       intros. unfold ehom. unfold comp. simpl. unfold cutProof.
112       rewrite <- (@ndr_prod_preserves_comp _ _ pl_eqv _ _ (cnd_expand_left _ _ c) _ _ (nd_id1 (b|=c0))
113                   _ (nd_id1 (c,,a |= c,,b))  _ (cnd_expand_left _ _ c)).
114       setoid_rewrite (@ndr_comp_right_identity _ _ pl_eqv _ [c,,a |= c,,b]).
115       setoid_rewrite (@ndr_comp_left_identity  _ _ pl_eqv [b |= c0]).
116       simpl; eapply cndr_inert. apply pl_eqv. auto. auto.
117       Defined.
118
119     Definition Types_binoidal : EBinoidalCat TypesL.
120       refine
121         {| ebc_first  := Types_first
122          ; ebc_second := Types_second
123          |}.
124       Defined.
125
126     Instance Types_assoc_iso a b c : Isomorphic(C:=TypesL) ((a,,b),,c) (a,,(b,,c)) :=
127       { iso_forward  := snd_initial _ ;; cnd_ant_cossa _ a b c
128       ; iso_backward := snd_initial _ ;; cnd_ant_assoc _ a b c
129       }.
130       simpl; eapply cndr_inert. unfold identityProof; apply pl_eqv. auto.
131         apply ndpc_comp; auto.
132         apply ndpc_comp; auto.
133         auto.
134       simpl; eapply cndr_inert. unfold identityProof; apply pl_eqv. auto.
135         apply ndpc_comp; auto.
136         apply ndpc_comp; auto.
137         auto.
138         Defined.
139
140     Instance Types_cancelr_iso a : Isomorphic(C:=TypesL) (a,,[]) a :=
141       { iso_forward  := snd_initial _ ;; cnd_ant_rlecnac _ a
142       ; iso_backward := snd_initial _ ;; cnd_ant_cancelr _ a
143       }.
144       unfold eqv; unfold comp; simpl.
145       eapply cndr_inert. apply pl_eqv. auto.
146         apply ndpc_comp; auto.
147         apply ndpc_comp; auto.
148         auto.
149       unfold eqv; unfold comp; simpl.
150       eapply cndr_inert. apply pl_eqv. auto.
151         apply ndpc_comp; auto.
152         apply ndpc_comp; auto.
153         auto.
154       Defined.
155
156     Instance Types_cancell_iso a : Isomorphic(C:=TypesL) ([],,a) a :=
157       { iso_forward  := snd_initial _ ;; cnd_ant_llecnac _ a
158       ; iso_backward := snd_initial _ ;; cnd_ant_cancell _ a
159       }.
160       unfold eqv; unfold comp; simpl.
161       eapply cndr_inert. apply pl_eqv. auto.
162         apply ndpc_comp; auto.
163         apply ndpc_comp; auto.
164         auto.
165       unfold eqv; unfold comp; simpl.
166       eapply cndr_inert. apply pl_eqv. auto.
167         apply ndpc_comp; auto.
168         apply ndpc_comp; auto.
169         auto.
170       Defined.
171
172     Instance Types_assoc a b : Types_second a >>>> Types_first b <~~~> Types_first b >>>> Types_second a :=
173       { ni_iso := fun c => Types_assoc_iso a c b }.
174 (*
175       intros. unfold functor_comp. unfold fmor.
176       Opaque Types_assoc_iso.
177       simpl.
178       
179       eapply cndr_inert.
180       apply pl_eqv.
181       apply ndpc_comp; auto.
182       apply ndpc_comp; auto.
183       apply ndpc_comp; auto.
184       apply ndpc_prod; auto.
185       apply ndpc_comp; auto.
186       apply ndpc_comp; auto.
187 *)
188 admit.
189       Defined.
190
191     Instance Types_cancelr   : Types_first [] <~~~> functor_id _ :=
192       { ni_iso := Types_cancelr_iso }.
193       intros; simpl.
194       admit.
195       Defined.
196
197     Instance Types_cancell   : Types_second [] <~~~> functor_id _ :=
198       { ni_iso := Types_cancell_iso }.
199       admit.
200       Defined.
201
202     Instance Types_assoc_ll a b : Types_second (a,,b) <~~~> Types_second b >>>> Types_second a :=
203       { ni_iso := fun c => Types_assoc_iso a b c }.
204       admit.
205       Defined.
206
207     Instance Types_assoc_rr a b : Types_first (a,,b) <~~~> Types_first a >>>> Types_first b :=
208       { ni_iso := fun c => iso_inv _ _ (Types_assoc_iso c a b) }.
209       admit.
210       Defined.
211
212     Instance TypesL_PreMonoidal : PreMonoidalCat Types_binoidal [] :=
213       { pmon_assoc    := Types_assoc
214       ; pmon_cancell  := Types_cancell
215       ; pmon_cancelr  := Types_cancelr
216       ; pmon_assoc_rr := Types_assoc_rr
217       ; pmon_assoc_ll := Types_assoc_ll
218       }.
219 (*
220       apply Build_Pentagon.
221         intros; simpl.
222         eapply cndr_inert. apply pl_eqv.
223         apply ndpc_comp.
224         apply ndpc_comp.
225         auto.
226         apply ndpc_comp.
227         apply ndpc_prod.
228         apply ndpc_comp.
229         apply ndpc_comp.
230         auto.
231         apply ndpc_comp.
232         auto.
233         auto.
234         auto.
235         auto.
236         auto.
237         auto.
238         apply ndpc_comp.
239         apply ndpc_comp.
240         auto.
241         apply ndpc_comp.
242         auto.
243         auto.
244         auto.
245       apply Build_Triangle; intros; simpl.
246         eapply cndr_inert. apply pl_eqv.
247         auto.
248         apply ndpc_comp.
249         apply ndpc_comp.
250         auto.
251         apply ndpc_comp.
252         auto.
253         auto.
254         auto.
255         eapply cndr_inert. apply pl_eqv. auto.
256           auto.
257 *)
258 admit.
259 admit.
260       intros; simpl; reflexivity.
261       intros; simpl; reflexivity.
262       admit.  (* assoc   central *)
263       admit.  (* cancelr central *)
264       admit.  (* cancell central *)
265       Defined.
266
267     Definition TypesEnrichedInJudgments : SurjectiveEnrichment.
268       refine
269         {| senr_c_pm     := TypesL_PreMonoidal
270          ; senr_v        := JudgmentsL
271          ; senr_v_bin    := Judgments_Category_binoidal _
272          ; senr_v_pmon   := Judgments_Category_premonoidal _
273          ; senr_v_mon    := Judgments_Category_monoidal _
274          ; senr_c_bin    := Types_binoidal
275          ; senr_c        := TypesL
276         |}.
277       Defined.
278
279   End LanguageCategory.
280
281 End Programming_Language.
282 Implicit Arguments ND [ Judgment ].