unbreak lots more stuff
[coq-hetmet.git] / src / HaskProofFlattener.v
index 980697d..7b70e6e 100644 (file)
@@ -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.
 
@@ -114,13 +116,12 @@ Section HaskProofFlattener.
     admit.
     Defined.
 
-  Definition guestJudgmentAsGArrowType {Γ}{Δ}{ec}(lt:PCFJudg Γ Δ ec) : HaskType Γ ★ :=
+  Definition guestJudgmentAsGArrowType {Γ}{Δ}(lt:PCFJudg Γ Δ) : HaskType Γ ★ :=
     match lt with
-      pcfjudg x y =>
-      (mkGA (tensorizeType x) (tensorizeType y)) 
+      (x,y) => (mkGA (tensorizeType x) (tensorizeType y)) 
     end.
 
-  Definition obact {Γ}{Δ} ec (X:Tree ??(PCFJudg Γ Δ ec)) : Tree ??(LeveledHaskType Γ ★) :=
+  Definition obact {Γ}{Δ} (X:Tree ??(PCFJudg Γ Δ)) : Tree ??(LeveledHaskType Γ ★) :=
     mapOptionTree guestJudgmentAsGArrowType X @@@ nil.
 
   Hint Constructors Rule_Flat.
@@ -131,94 +132,107 @@ Section HaskProofFlattener.
    * it might help to have the definition for "Inductive ND" (see
    * NaturalDeduction.v) handy as a cross-reference.
    *)
+  Hint Constructors Rule_Flat.
   Definition FlatteningFunctor_fmor {Γ}{Δ}{ec}
     : forall h c,
-      (h~~{JudgmentsL _ _ (PCF _ Γ Δ ec)}~~>c) ->
-      ((obact ec h)~~{TypesL _ _ (SystemFCa _ Γ Δ)}~~>(obact ec c)).
+      (h~~{JudgmentsL (PCF Γ Δ 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.
 
-    (* the proof from no hypotheses of no conclusions (nd_id0) becomes REmptyGroup *)
-    apply nd_rule; apply (org_fc _ _ (REmptyGroup _ _ )). auto.
+    (* the proof from no hypotheses of no conclusions (nd_id0) becomes RVoid *)
+    apply nd_rule; apply (org_fc _ _ [] [(_,_)] (RVoid _ _)). apply Flat_RVoid.
 
     (* the proof from hypothesis X of conclusion X (nd_id1) becomes RVar *)
-    apply nd_rule; apply (org_fc _ _ (RVar _ _ _ _)). auto.
+    apply nd_rule; apply (org_fc _ _ [] [(_,_)] (RVar _ _ _ _)). apply Flat_RVar.
 
-    (* the proof from hypothesis X of no conclusions (nd_weak) becomes RWeak;;REmptyGroup *)
+    (* the proof from hypothesis X of no conclusions (nd_weak) becomes RWeak;;RVoid *)
     eapply nd_comp;
       [ idtac
       | eapply nd_rule
-      ; eapply (org_fc  _ _ (RArrange _ _ _ _ _ (RWeak _)))
+      ; eapply (org_fc  _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RWeak _)))
       ; auto ].
       eapply nd_rule.
-      eapply (org_fc _ _ (REmptyGroup _ _)); auto.
-    
-    (* the proof from hypothesis X of two identical conclusions X,,X (nd_copy) becomes RVar;;RBindingGroup;;RCont *)
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply (org_fc _ _ (RArrange _ _ _ _ _ (RCont _))) ].
+      eapply (org_fc _ _ [] [(_,_)] (RVoid _ _)); auto. apply Flat_RVoid.
+      apply Flat_RArrange.
+
+    (* 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 _ Γ Δ))
-        (mapOptionTree guestJudgmentAsGArrowType h @@@ lev)) as q.
+      set (snd_initial(SequentND:=pl_snd(ProgrammingLanguage:=SystemFCa Γ Δ))
+        (mapOptionTree (guestJudgmentAsGArrowType(Δ:=ec)) h @@@ lev)) as q.
       eapply nd_comp.
       eapply nd_prod.
       apply q.
       apply q.
       apply nd_rule. 
-      eapply (org_fc _ _ (RBindingGroup _ _ _ _ _ _ )).
-      auto.
-      auto.
-
-    (* nd_prod becomes nd_llecnac;;nd_prod;;RBindingGroup *)
+      eapply (org_fc _ _ ([(_,_)],,[(_,_)]) [(_,_)] (RJoin _ _ _ _ _ _ )).
+      destruct h; simpl.
+      destruct o.
+      simpl.
+      apply Flat_RJoin.
+      apply Flat_RJoin.
+      apply Flat_RJoin.
+      apply Flat_RArrange.
+
+    (* nd_prod becomes nd_llecnac;;nd_prod;;RJoin *)
     eapply nd_comp.
       apply (nd_llecnac ;; nd_prod IHX1 IHX2).
       apply nd_rule.
-      eapply (org_fc _ _ (RBindingGroup _ _ _ _ _ _ )).
-      auto.
+      eapply (org_fc _ _ ([(_,_)],,[(_,_)]) [(_,_)] (RJoin _ _ _ _ _ _ )).
+      apply (Flat_RJoin Γ Δ (mapOptionTree guestJudgmentAsGArrowType h1 @@@ nil)
+       (mapOptionTree guestJudgmentAsGArrowType h2 @@@ nil)
+       (mapOptionTree guestJudgmentAsGArrowType c1 @@@ nil)
+       (mapOptionTree guestJudgmentAsGArrowType c2 @@@ nil)).
 
     (* nd_comp becomes pl_subst (aka nd_cut) *)
     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 _ Γ Δ))).
-      auto.
+      [ idtac | eapply nd_rule; apply (org_fc _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RuCanL _))) ].
+      apply (snd_initial(SequentND:=pl_cnd(ProgrammingLanguage:=(SystemFCa Γ Δ)))).
+      apply Flat_RArrange.
 
     (* 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 _ Γ Δ))).
-      auto.
+      [ idtac | eapply nd_rule; apply (org_fc _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RuCanR _))) ].
+      apply (snd_initial(SequentND:=pl_cnd(ProgrammingLanguage:=(SystemFCa Γ Δ)))).
+      apply Flat_RArrange.
 
     (* 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 _ Γ Δ))).
-      auto.
+      [ idtac | eapply nd_rule; apply (org_fc _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RCanL _))) ].
+      apply (snd_initial(SequentND:=pl_cnd(ProgrammingLanguage:=(SystemFCa Γ Δ)))).
+      apply Flat_RArrange.
 
     (* 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 _ Γ Δ))).
-      auto.
+      [ idtac | eapply nd_rule; apply (org_fc _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RCanR _))) ].
+      apply (snd_initial(SequentND:=pl_cnd(ProgrammingLanguage:=(SystemFCa Γ Δ)))).
+      apply Flat_RArrange.
 
     (* 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 _ Γ Δ))).
-      auto.
+      [ idtac | eapply nd_rule; apply (org_fc _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RAssoc _ _ _))) ].
+      apply (snd_initial(SequentND:=pl_cnd(ProgrammingLanguage:=(SystemFCa Γ Δ)))).
+      apply Flat_RArrange.
 
     (* 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 _ Γ Δ))).
-      auto.
+      [ idtac | eapply nd_rule; apply (org_fc _ _ [(_,_)] [(_,_)] (RArrange _ _ _ _ _ (RCossa _ _ _))) ].
+      apply (snd_initial(SequentND:=pl_cnd(ProgrammingLanguage:=(SystemFCa Γ Δ)))).
+      apply Flat_RArrange.
 
       destruct r as [r rp].
       refine (match rp as R in @Rule_PCF _ _ _ H C _ with
@@ -229,8 +243,8 @@ Section HaskProofFlattener.
                 | PCF_RLam             Σ tx te          => let case_RLam          := tt in _
                 | PCF_RApp             Σ tx te   p      => let case_RApp          := tt in _
                 | PCF_RLet             Σ σ₁ σ₂   p      => let case_RLet          := tt in _
-                | PCF_RBindingGroup    b c d e          => let case_RBindingGroup := tt in _
-                | PCF_REmptyGroup                       => let case_REmptyGroup   := tt in _
+                | PCF_RJoin    b c d e          => let case_RJoin := tt in _
+                | PCF_RVoid                       => let case_RVoid   := tt in _
               (*| PCF_RCase            T κlen κ θ l x   => let case_RCase         := tt in _*)
               (*| PCF_RLetRec          Σ₁ τ₁ τ₂ lev     => let case_RLetRec       := tt in _*)
               end); simpl in *.
@@ -310,9 +324,11 @@ Section HaskProofFlattener.
       destruct case_RNote.
         eapply nd_comp.
         eapply nd_rule.
-        eapply (org_fc _ _ (RVar _ _ _ _)) . auto.
+        eapply (org_fc _ _ [] [(_,_)] (RVar _ _ _ _)) . auto.
+        apply Flat_RVar.
         apply nd_rule.
-        apply (org_fc _ _ (RNote _ _ _ _ _ n)). auto.
+        apply (org_fc _ _ [(_,_)] [(_,_)] (RNote _ _ _ _ _ n)). auto.
+        apply Flat_RNote.
         
       destruct case_RVar.
         (* ga_id *)
@@ -330,33 +346,33 @@ Section HaskProofFlattener.
         (* ga_comp! perhaps this means the ga_curry avoidance can be done by turning lambdas into lets? *)
         admit.
 
-      destruct case_REmptyGroup.
+      destruct case_RVoid.
         (* ga_id u *)
         admit.
 
-      destruct case_RBindingGroup.
+      destruct case_RJoin.
         (* ga_first+ga_second; technically this assumes a specific evaluation order, which is bad *)
         admit.
 
       Defined.
 
-  Instance FlatteningFunctor {Γ}{Δ}{ec} : Functor (JudgmentsL _ _ (PCF _ Γ Δ ec)) (TypesL _ _ (SystemFCa _ Γ Δ)) (obact ec) :=
+  Instance FlatteningFunctor {Γ}{Δ}{ec} : Functor (JudgmentsL (PCF Γ Δ ec)) (TypesL (SystemFCa Γ Δ)) (obact) :=
     { fmor := FlatteningFunctor_fmor }.
     admit.
     admit.
     admit.
     Defined.
-(*
-    Definition ReificationFunctor Γ Δ : Functor (JudgmentsL _ _ (PCF n Γ Δ)) SystemFCa' (mapOptionTree brakifyJudg).
-      refine {| fmor := ReificationFunctor_fmor Γ Δ |}; unfold hom; unfold ob; simpl ; intros.
-      unfold ReificationFunctor_fmor; simpl.
-      admit.
-      unfold ReificationFunctor_fmor; simpl.
-      admit.
-      unfold ReificationFunctor_fmor; simpl.
-      admit.
-      Defined.
 
+  (*
+  Definition ReificationFunctor Γ Δ : Functor (JudgmentsL _ _ (PCF n Γ Δ)) SystemFCa' (mapOptionTree brakifyJudg).
+    refine {| fmor := ReificationFunctor_fmor Γ Δ |}; unfold hom; unfold ob; simpl ; intros.
+    unfold ReificationFunctor_fmor; simpl.
+    admit.
+    unfold ReificationFunctor_fmor; simpl.
+    admit.
+    unfold ReificationFunctor_fmor; simpl.
+    admit.
+    Defined.
 
   Definition PCF_SMME (n:nat)(Γ:TypeEnv)(Δ:CoercionEnv Γ) : ProgrammingLanguageSMME.
     refine {| plsmme_pl := PCF n Γ Δ |}.
@@ -377,7 +393,7 @@ Section HaskProofFlattener.
     admit.
     (*  ... and the retraction exists *)
     Defined.
-*)
+  *)
   (* Any particular proof in HaskProof is only finitely large, so it uses only finitely many levels of nesting, so
    * it falls within (SystemFCa n) for some n.  This function calculates that "n" and performs the translation *)
   (*
@@ -385,7 +401,6 @@ Section HaskProofFlattener.
     forall h c (pf:ND Rule h c),
       { n:nat & h ~~{JudgmentsL (SystemFCa_SMME n)}~~> c }.
       *)
-
   (* for every n we have a functor from the category of (n+1)-bounded proofs to the category of n-bounded proofs *)
 
 End HaskProofFlattener.