X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskProofFlattener.v;h=fcc2a37a78d865fe4500131e4df63eb888fc72bc;hp=c37079f2686eda6e4f773c4c855cc79dfa4d04a1;hb=6c953e094065d487e85635df7fd5389271e4a279;hpb=6ef9f270b138fc7aab48013d55a8192ff022c0f1 diff --git a/src/HaskProofFlattener.v b/src/HaskProofFlattener.v index c37079f..fcc2a37 100644 --- a/src/HaskProofFlattener.v +++ b/src/HaskProofFlattener.v @@ -30,6 +30,8 @@ Require Import Enrichment_ch2_8. Require Import Subcategories_ch7_1. Require Import NaturalTransformations_ch7_4. Require Import NaturalIsomorphisms_ch7_5. +Require Import BinoidalCategories. +Require Import PreMonoidalCategories. Require Import MonoidalCategories_ch7_8. Require Import Coherence_ch7_8. @@ -131,14 +133,15 @@ Section HaskProofFlattener. * it might help to have the definition for "Inductive ND" (see * NaturalDeduction.v) handy as a cross-reference. *) +(* Definition FlatteningFunctor_fmor {Γ}{Δ}{ec} : forall h c, (h~~{JudgmentsL _ _ (PCF _ Γ Δ ec)}~~>c) -> - ((obact ec h)~~{TypesL _ _ (SystemFCa _ Γ Δ)}~~>(obact ec c)). + ((obact ec h)~~{TypesL _ _ (SystemFCa Γ Δ)}~~>(obact ec c)). set (@nil (HaskTyVar Γ ★)) as lev. - unfold hom; unfold ob; unfold ehom; simpl; unfold mon_i; unfold obact; intros. + unfold hom; unfold ob; unfold ehom; simpl; unfold pmon_I; unfold obact; intros. induction X; simpl. @@ -160,7 +163,7 @@ Section HaskProofFlattener. (* the proof from hypothesis X of two identical conclusions X,,X (nd_copy) becomes RVar;;RJoin;;RCont *) eapply nd_comp; [ idtac | eapply nd_rule; eapply (org_fc _ _ (RArrange _ _ _ _ _ (RCont _))) ]. eapply nd_comp; [ apply nd_llecnac | idtac ]. - set (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ)) + set (snd_initial(SequentND:=@pl_snd _ _ _ _ (SystemFCa Γ Δ)) (mapOptionTree guestJudgmentAsGArrowType h @@@ lev)) as q. eapply nd_comp. eapply nd_prod. @@ -182,42 +185,45 @@ Section HaskProofFlattener. eapply nd_comp. apply (nd_llecnac ;; nd_prod IHX1 IHX2). clear IHX1 IHX2 X1 X2. - apply (@nd_cut _ _ _ _ _ _ (@pl_subst _ _ _ _ (SystemFCa _ Γ Δ))). + (* + apply (@snd_cut _ _ _ _ _ _ (@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). + *) + admit. (* nd_cancell becomes RVar;;RuCanL *) eapply nd_comp; [ idtac | eapply nd_rule; apply (org_fc _ _ (RArrange _ _ _ _ _ (RuCanL _))) ]. - apply (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ))). + apply (snd_initial(SequentND:=@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). auto. (* nd_cancelr becomes RVar;;RuCanR *) eapply nd_comp; [ idtac | eapply nd_rule; apply (org_fc _ _ (RArrange _ _ _ _ _ (RuCanR _))) ]. - apply (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ))). + apply (snd_initial(SequentND:=@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). auto. (* nd_llecnac becomes RVar;;RCanL *) eapply nd_comp; [ idtac | eapply nd_rule; apply (org_fc _ _ (RArrange _ _ _ _ _ (RCanL _))) ]. - apply (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ))). + apply (snd_initial(SequentND:=@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). auto. (* nd_rlecnac becomes RVar;;RCanR *) eapply nd_comp; [ idtac | eapply nd_rule; apply (org_fc _ _ (RArrange _ _ _ _ _ (RCanR _))) ]. - apply (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ))). + apply (snd_initial(SequentND:=@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). auto. (* nd_assoc becomes RVar;;RAssoc *) eapply nd_comp; [ idtac | eapply nd_rule; apply (org_fc _ _ (RArrange _ _ _ _ _ (RAssoc _ _ _))) ]. - apply (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ))). + apply (snd_initial(SequentND:=@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). auto. (* nd_cossa becomes RVar;;RCossa *) eapply nd_comp; [ idtac | eapply nd_rule; apply (org_fc _ _ (RArrange _ _ _ _ _ (RCossa _ _ _))) ]. - apply (nd_seq_reflexive(SequentCalculus:=@pl_sc _ _ _ _ (SystemFCa _ Γ Δ))). + apply (snd_initial(SequentND:=@pl_cnd _ _ _ _ (SystemFCa Γ Δ))). auto. destruct r as [r rp]. @@ -340,7 +346,7 @@ Section HaskProofFlattener. Defined. - Instance FlatteningFunctor {Γ}{Δ}{ec} : Functor (JudgmentsL _ _ (PCF _ Γ Δ ec)) (TypesL _ _ (SystemFCa _ Γ Δ)) (obact ec) := + Instance FlatteningFunctor {Γ}{Δ}{ec} : Functor (JudgmentsL _ _ (PCF _ Γ Δ ec)) (TypesL _ _ (SystemFCa Γ Δ)) (obact ec) := { fmor := FlatteningFunctor_fmor }. admit. admit. @@ -387,6 +393,6 @@ Section HaskProofFlattener. *) (* for every n we have a functor from the category of (n+1)-bounded proofs to the category of n-bounded proofs *) - +*) End HaskProofFlattener.