Merge branch 'coq-extraction-baked-in' of /afs/megacz.com/.pub/software/coq-hetmet
authorAdam Megacz <megacz@gentzen.megacz.com>
Mon, 30 May 2011 14:11:57 +0000 (07:11 -0700)
committerAdam Megacz <megacz@gentzen.megacz.com>
Mon, 30 May 2011 14:11:57 +0000 (07:11 -0700)
17 files changed:
examples/Demo.hs
examples/Makefile
src/ExtractionMain.v
src/General.v
src/HaskFlattener.v
src/HaskProgrammingLanguage.v
src/HaskProof.v
src/HaskProofToLatex.v
src/HaskProofToStrong.v
src/HaskSkolemizer.v
src/HaskStrong.v
src/HaskStrongToProof.v
src/HaskStrongToWeak.v
src/HaskWeak.v
src/HaskWeakToStrong.v
src/NaturalDeductionContext.v [new file with mode: 0644]
src/PCF.v

index 4c53aa5..6d0e69d 100644 (file)
@@ -1,4 +1,4 @@
-{-# OPTIONS_GHC -XModalTypes -fcoqpass -dcore-lint #-}
+{-# OPTIONS_GHC -XModalTypes -fflatten -funsafe-skolemize -dcore-lint #-}
 module Demo (demo) where
 
 
index c89fb10..b9fa6a7 100644 (file)
@@ -4,11 +4,23 @@ open:
        make demo
        open .build/test.pdf
 
-all:
+#sanity += BiGArrow.hs
+#sanity += CircuitExample.hs
+sanity += CommandSyntaxExample.hs
+sanity += DotProduct.hs
+sanity += GArrowTutorial.hs
+sanity += GArrowVerilog.hs
+sanity += ImmutableHeap.hs
+sanity += IsomorphismForCodeTypes.hs
+sanity += LambdaCalculusInterpreter.hs
+sanity += TypeSafeRun.hs
+sanity += Unflattening.hs
+
+sanity:
        ../../../inplace/bin/ghc-stage2 -dcore-lint -fforce-recomp -fcoqpass -ddump-coqpass -ddump-to-file \
-               `ls *.hs | grep -v Regex | grep -v Unify.hs | grep -v GArrowTikZ.hs ` +RTS -K500M 
-       ../../../inplace/bin/ghc-stage2 -dcore-lint -fforce-recomp \
-               RegexMatcher.hs Unify.hs GArrowTikZ.hs
+               $(sanity) \
+               +RTS -K500M 
+
 
 demo:
        mkdir -p .build
index 77326e0..4926bff 100644 (file)
@@ -13,6 +13,7 @@ Require Import Preamble.
 Require Import General.
 
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 
 Require Import HaskKinds.
 Require Import HaskLiterals.
@@ -132,7 +133,7 @@ Section core2proof.
                     OK (eol+++eol+++eol+++
                         "\begin{preview}"+++eol+++
                         "$\displaystyle "+++
-                        toString (nd_ml_toLatexMath (@expr2proof _ _ _ _ _ _ e))+++
+                        toString (nd_ml_toLatexMath (@expr2proof _ _ _ _ _ _ _ e))+++
                         " $"+++eol+++
                         "\end{preview}"+++eol+++eol+++eol)
                   )))))))).
@@ -236,7 +237,7 @@ Section core2proof.
     ND Rule 
        [ Γ > Δ > Σ             |- [a ---> s  ]@lev ]
        [ Γ > Δ > [a @@ lev],,Σ |-       [ s ]@lev ].
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     eapply nd_comp; [ idtac | eapply nd_rule; eapply RApp ].
     eapply nd_comp; [ apply nd_rlecnac | idtac ].
     apply nd_prod.
@@ -251,7 +252,7 @@ Section core2proof.
     intro pf.
     eapply nd_comp.
     apply pf.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply ACanR ].
     apply curry.
     Defined.
 
@@ -267,13 +268,15 @@ Section core2proof.
     eapply nd_comp.
     eapply nd_rule.
     eapply RArrange.
-    eapply RCanR.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply ACanR.
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     apply curry.
     Defined.
 
   Section coqPassCoreToCore.
     Context
+    (do_flatten : bool)
+    (do_skolemize : bool)
     (hetmet_brak  : CoreVar)
     (hetmet_esc   : CoreVar)
     (hetmet_flatten : CoreVar)
@@ -431,18 +434,41 @@ Section core2proof.
                     ((weakExprToStrongExpr Γ Δ φ ψ ξ (fun _ => true) τ nil we) >>= fun e =>
 
                       (addErrorMessage ("HaskStrong...")
-                        (let haskProof := skolemize_and_flatten_proof hetmet_flatten' hetmet_unflatten'
-                                            hetmet_flattened_id' my_ga (@expr2proof _ _ _ _ _ _ e)
-                         in (* insert HaskProof-to-HaskProof manipulations here *)
-                         OK ((@proof2expr nat _ FreshNat _ _ (flatten_type τ@@nil) _ (fun _ => Prelude_error "unbound unique") _ haskProof) O)
-                         >>= fun e' =>
-                           (snd e') >>= fun e'' =>
-                         strongExprToWeakExpr hetmet_brak' hetmet_esc'
-                           mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
-                           (projT2 e'') INil
-                         >>= fun q =>
-                           OK (weakExprToCoreExpr q)
-                    )))))))))).
+                        (if do_skolemize
+                        then
+                             (let haskProof := skolemize_and_flatten_proof hetmet_flatten' hetmet_unflatten'
+                                                 hetmet_flattened_id' my_ga (@expr2proof _ _ _ _ _ _ _ e)
+                              in (* insert HaskProof-to-HaskProof manipulations here *)
+                              OK ((@proof2expr nat _ FreshNat _ _ (flatten_type τ) nil _
+                                (fun _ => Prelude_error "unbound unique") _ haskProof) O)
+                              >>= fun e' => (snd e') >>= fun e'' =>
+                              strongExprToWeakExpr hetmet_brak' hetmet_esc'
+                                mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
+                                (projT2 e'') INil
+                              >>= fun q => OK (weakExprToCoreExpr q))
+                        else (if do_flatten
+                        then
+                          (let haskProof := flatten_proof (*hetmet_flatten' hetmet_unflatten'
+                                                 hetmet_flattened_id' my_ga*) (@expr2proof _ _ _ _ _ _ _ e)
+                              in (* insert HaskProof-to-HaskProof manipulations here *)
+                              OK ((@proof2expr nat _ FreshNat _ _ τ nil _
+                                (fun _ => Prelude_error "unbound unique") _ haskProof) O)
+                              >>= fun e' => (snd e') >>= fun e'' =>
+                              strongExprToWeakExpr hetmet_brak' hetmet_esc'
+                                mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
+                                (projT2 e'') INil
+                              >>= fun q => OK (weakExprToCoreExpr q))
+                        else
+                          (let haskProof := @expr2proof _ _ _ _ _ _ _ e
+                              in (* insert HaskProof-to-HaskProof manipulations here *)
+                              OK ((@proof2expr nat _ FreshNat _ _ τ nil _
+                                (fun _ => Prelude_error "unbound unique") _ haskProof) O)
+                              >>= fun e' => (snd e') >>= fun e'' =>
+                              strongExprToWeakExpr hetmet_brak' hetmet_esc'
+                                mkWeakTypeVar mkWeakCoerVar mkWeakExprVar uniqueSupply
+                                (projT2 e'') INil
+                              >>= fun q => OK (weakExprToCoreExpr q))))
+                  ))))))))).
 
     Definition coreToCoreExpr (ce:@CoreExpr CoreVar) : (@CoreExpr CoreVar) :=
       match coreToCoreExpr' ce with
@@ -472,6 +498,8 @@ Section core2proof.
   End coqPassCoreToCore.
 
     Definition coqPassCoreToCore 
+    (do_flatten  : bool)
+    (do_skolemize  : bool)
     (hetmet_brak  : CoreVar)
     (hetmet_esc   : CoreVar)
     (hetmet_flatten   : CoreVar)
@@ -501,6 +529,8 @@ Section core2proof.
     (hetmet_pga_curryl : CoreVar)
     (hetmet_pga_curryr : CoreVar) : list (@CoreBind CoreVar) :=
     coqPassCoreToCore'
+       do_flatten
+       do_skolemize
        hetmet_brak  
        hetmet_esc   
        hetmet_flatten
index ae27b9f..ee5cb16 100644 (file)
@@ -106,6 +106,11 @@ Fixpoint treeReduce {T:Type}{R:Type}(mapLeaf:T->R)(mergeBranches:R->R->R) (t:Tre
   end.
 Definition treeDecomposition {D T:Type} (mapLeaf:T->D) (mergeBranches:D->D->D) :=
   forall d:D, { tt:Tree T & d = treeReduce mapLeaf mergeBranches tt }.
+Lemma mapOptionTree_distributes
+  : forall T R (a b:Tree ??T) (f:T->R),
+    mapOptionTree f (a,,b) = (mapOptionTree f a),,(mapOptionTree f b).
+  reflexivity.
+  Qed.
 
 Fixpoint reduceTree {T}(unit:T)(merge:T -> T -> T)(tt:Tree ??T) : T :=
   match tt with
@@ -520,6 +525,12 @@ Fixpoint takeT {T}{Σ:Tree T}(tf:TreeFlags Σ) : ??(Tree T) :=
       end
   end.
 
+Definition takeT' {T}{Σ:Tree ??T}(tf:TreeFlags Σ) : Tree ??T :=
+  match takeT tf with
+    | None   => []
+    | Some x => x
+  end.
+
 (* lift a function T->bool to a function (option T)->bool by yielding (None |-> b) *)
 Definition liftBoolFunc {T}(b:bool)(f:T -> bool) : ??T -> bool :=
   fun t =>
index 9b3a163..d805de4 100644 (file)
@@ -9,6 +9,7 @@ Generalizable All Variables.
 Require Import Preamble.
 Require Import General.
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 Require Import Coq.Strings.String.
 Require Import Coq.Lists.List.
 
@@ -45,52 +46,6 @@ Set Printing Width 130.
  *)
 Section HaskFlattener.
 
-  Definition getlev {Γ}{κ}(lht:LeveledHaskType Γ κ) : HaskLevel Γ :=
-    match lht with t @@ l => l end.
-
-  Definition arrange :
-    forall {T} (Σ:Tree ??T) (f:T -> bool),
-      Arrange Σ (dropT (mkFlags (liftBoolFunc false f) Σ),,( (dropT (mkFlags (liftBoolFunc false (bnot ○ f)) Σ)))).
-    intros.
-    induction Σ.
-      simpl.
-      destruct a.
-      simpl.
-      destruct (f t); simpl.
-      apply RuCanL.
-      apply RuCanR.
-      simpl.
-      apply RuCanL.
-      simpl in *.
-      eapply RComp; [ idtac | apply arrangeSwapMiddle ].
-      eapply RComp.
-      eapply RLeft.
-      apply IHΣ2.
-      eapply RRight.
-      apply IHΣ1.
-      Defined.
-
-  Definition arrange' :
-    forall {T} (Σ:Tree ??T) (f:T -> bool),
-      Arrange (dropT (mkFlags (liftBoolFunc false f) Σ),,( (dropT (mkFlags (liftBoolFunc false (bnot ○ f)) Σ)))) Σ.
-    intros.
-    induction Σ.
-      simpl.
-      destruct a.
-      simpl.
-      destruct (f t); simpl.
-      apply RCanL.
-      apply RCanR.
-      simpl.
-      apply RCanL.
-      simpl in *.
-      eapply RComp; [ apply arrangeSwapMiddle | idtac ].
-      eapply RComp.
-      eapply RLeft.
-      apply IHΣ2.
-      eapply RRight.
-      apply IHΣ1.
-      Defined.
 
   Ltac eqd_dec_refl' :=
     match goal with
@@ -177,29 +132,11 @@ Section HaskFlattener.
     rewrite <- IHx2 at 2.
     reflexivity.
     Qed.
-(*
-  Lemma drop_lev_lemma' : forall Γ (lev:HaskLevel Γ) x, drop_lev lev (x @@@ lev) = [].
-    intros.
-    induction x.
-    destruct a; simpl; try reflexivity.
-    apply drop_lev_lemma.
-    simpl.
-    change (@drop_lev _ lev (x1 @@@ lev ,, x2 @@@ lev))
-      with ((@drop_lev _ lev (x1 @@@ lev)) ,, (@drop_lev _ lev (x2 @@@ lev))).
-    simpl.
-    rewrite IHx1.
-    rewrite IHx2.
-    reflexivity.
-    Qed.
-*)
+
   Ltac drop_simplify :=
     match goal with
       | [ |- context[@drop_lev ?G ?L [ ?X @@ ?L ] ] ] =>
         rewrite (drop_lev_lemma G L X)
-(*
-      | [ |- context[@drop_lev ?G ?L [ ?X @@@ ?L ] ] ] =>
-        rewrite (drop_lev_lemma' G L X)
-*)
       | [ |- context[@drop_lev ?G (?E :: ?L) [ ?X @@ (?E :: ?L) ] ] ] =>
         rewrite (drop_lev_lemma_s G L E X)
       | [ |- context[@drop_lev ?G ?N (?A,,?B)] ] =>
@@ -350,7 +287,7 @@ Section HaskFlattener.
     ND Rule []                         [ Γ > Δ > [x@@lev] |- [y]@lev ] ->
     ND Rule [ Γ > Δ > ant |- [x]@lev ] [ Γ > Δ > ant      |- [y]@lev ].
     intros.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanR ].
     eapply nd_comp; [ idtac | eapply nd_rule; apply RLet ].
     eapply nd_comp; [ apply nd_rlecnac | idtac ].
     apply nd_prod.
@@ -359,7 +296,7 @@ Section HaskFlattener.
       apply X.
       eapply nd_rule.
       eapply RArrange.
-      apply RuCanR.
+      apply AuCanR.
     Defined.
 
   Definition precompose Γ Δ ec : forall a x y z lev,
@@ -371,7 +308,7 @@ Section HaskFlattener.
     eapply nd_comp; [ apply nd_rlecnac | idtac ].
     apply nd_prod.
     apply nd_id.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     apply ga_comp.
     Defined.
 
@@ -380,8 +317,8 @@ Section HaskFlattener.
       [ Γ > Δ > a,,b                             |- [@ga_mk _ ec y z ]@lev ]
       [ Γ > Δ > a,,([@ga_mk _ ec x y @@ lev],,b) |- [@ga_mk _ ec x z ]@lev ].
     intros.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RLeft; eapply RExch ].
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCossa ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ALeft; eapply AExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AuAssoc ].
     apply precompose.
     Defined.
 
@@ -401,7 +338,7 @@ Section HaskFlattener.
     ND Rule [] [ Γ > Δ > []                       |- [@ga_mk _ ec x y ]@lev ] ->
     ND Rule [] [ Γ > Δ > [@ga_mk _ ec y z @@ lev] |- [@ga_mk _ ec x z ]@lev ].
     intros.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanL ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanL ].
     eapply nd_comp; [ idtac | eapply postcompose_ ].
     apply X.
     Defined.
@@ -410,12 +347,12 @@ Section HaskFlattener.
     ND Rule [ Γ > Δ > Σ                    |- [@ga_mk Γ ec a b ]@lev ]
             [ Γ > Δ > Σ                    |- [@ga_mk Γ ec (a,,c) (b,,c) ]@lev ].
     intros.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanR ].
     eapply nd_comp; [ idtac | eapply nd_rule; apply RLet ].
     eapply nd_comp; [ apply nd_rlecnac | idtac ].
     apply nd_prod.
     apply nd_id.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RuCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AuCanR ].
     apply ga_first.
     Defined.
 
@@ -433,12 +370,12 @@ Section HaskFlattener.
      [ Γ > Δ > Σ                    |- [@ga_mk Γ ec a b ]@lev ]
      [ Γ > Δ > Σ                    |- [@ga_mk Γ ec (c,,a) (c,,b) ]@lev ].
     intros.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanR ].
     eapply nd_comp; [ idtac | eapply nd_rule; apply RLet ].
     eapply nd_comp; [ apply nd_rlecnac | idtac ].
     apply nd_prod.
     apply nd_id.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RuCanR ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AuCanR ].
     apply ga_second.
     Defined.
 
@@ -456,7 +393,7 @@ Section HaskFlattener.
      [Γ > Δ > Σ                          |- [@ga_mk Γ ec (a,,x)  b ]@l ] 
      [Γ > Δ > Σ,,[@ga_mk Γ ec [] a @@ l] |- [@ga_mk Γ ec x       b ]@l ].
      intros.
-     eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+     eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
      eapply nd_comp; [ idtac | eapply nd_rule; eapply RLet ].
      eapply nd_comp; [ apply nd_llecnac | idtac ].
      apply nd_prod.
@@ -467,7 +404,7 @@ Section HaskFlattener.
      apply nd_prod.
      apply postcompose.
      apply ga_uncancell.
-     eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+     eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
      apply precompose.
      Defined.
 
@@ -493,38 +430,38 @@ Section HaskFlattener.
             (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) B))
             (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) A)) ]@nil]
           with
-          | RId  a               => let case_RId := tt    in ga_id _ _ _ _ _
-          | RCanL  a             => let case_RCanL := tt  in ga_uncancell _ _ _ _ _
-          | RCanR  a             => let case_RCanR := tt  in ga_uncancelr _ _ _ _ _
-          | RuCanL a             => let case_RuCanL := tt in ga_cancell _ _ _ _ _
-          | RuCanR a             => let case_RuCanR := tt in ga_cancelr _ _ _ _ _
-          | RAssoc a b c         => let case_RAssoc := tt in ga_assoc _ _ _ _ _ _ _
-          | RCossa a b c         => let case_RCossa := tt in ga_unassoc _ _ _ _ _ _ _
-          | RExch  a b           => let case_RExch := tt  in ga_swap  _ _ _ _ _ _
-          | RWeak  a             => let case_RWeak := tt  in ga_drop _ _ _ _ _ 
-          | RCont  a             => let case_RCont := tt  in ga_copy  _ _ _ _ _ 
-          | RLeft  a b c r'      => let case_RLeft := tt  in flatten _ _ r' ;; boost _ _ _ _ _ (ga_second _ _ _ _ _ _ _)
-          | RRight a b c r'      => let case_RRight := tt in flatten _ _ r' ;; boost _ _ _ _ _ (ga_first  _ _ _ _ _ _ _)
-          | RComp  c b a r1 r2   => let case_RComp := tt  in (fun r1' r2' => _) (flatten _ _ r1) (flatten _ _ r2)
+          | AId  a               => let case_AId := tt    in ga_id _ _ _ _ _
+          | ACanL  a             => let case_ACanL := tt  in ga_uncancell _ _ _ _ _
+          | ACanR  a             => let case_ACanR := tt  in ga_uncancelr _ _ _ _ _
+          | AuCanL a             => let case_AuCanL := tt in ga_cancell _ _ _ _ _
+          | AuCanR a             => let case_AuCanR := tt in ga_cancelr _ _ _ _ _
+          | AAssoc a b c         => let case_AAssoc := tt in ga_assoc _ _ _ _ _ _ _
+          | AuAssoc a b c         => let case_AuAssoc := tt in ga_unassoc _ _ _ _ _ _ _
+          | AExch  a b           => let case_AExch := tt  in ga_swap  _ _ _ _ _ _
+          | AWeak  a             => let case_AWeak := tt  in ga_drop _ _ _ _ _ 
+          | ACont  a             => let case_ACont := tt  in ga_copy  _ _ _ _ _ 
+          | ALeft  a b c r'      => let case_ALeft := tt  in flatten _ _ r' ;; boost _ _ _ _ _ (ga_second _ _ _ _ _ _ _)
+          | ARight a b c r'      => let case_ARight := tt in flatten _ _ r' ;; boost _ _ _ _ _ (ga_first  _ _ _ _ _ _ _)
+          | AComp  c b a r1 r2   => let case_AComp := tt  in (fun r1' r2' => _) (flatten _ _ r1) (flatten _ _ r2)
         end); clear flatten; repeat take_simplify; repeat drop_simplify; intros.
 
-        destruct case_RComp.
+        destruct case_AComp.
           set (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) a)) as a' in *.
           set (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) b)) as b' in *.
           set (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) c)) as c' in *.
-          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RCanL ].
+          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply ACanL ].
           eapply nd_comp; [ idtac | eapply nd_rule; apply
              (@RLet Γ Δ [] [] (@ga_mk _ (v2t ec) a' b') (@ga_mk _ (v2t ec) a' c')) ].
           eapply nd_comp; [ apply nd_llecnac | idtac ].
           apply nd_prod.
           apply r2'.
-          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RuCanR ].
-          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RCanL ].
+          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply AuCanR ].
+          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply ACanL ].
           eapply nd_comp; [ idtac | eapply nd_rule;  apply RLet ].
           eapply nd_comp; [ apply nd_llecnac | idtac ].
           eapply nd_prod.
           apply r1'.
-          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+          eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
           apply ga_comp.
           Defined.
 
@@ -548,19 +485,19 @@ Section HaskFlattener.
         match r as R in Arrange A B return
           Arrange (mapOptionTree (flatten_leveled_type ) (drop_lev _ A))
           (mapOptionTree (flatten_leveled_type ) (drop_lev _ B)) with
-          | RId  a               => let case_RId := tt  in RId _
-          | RCanL  a             => let case_RCanL := tt  in RCanL _
-          | RCanR  a             => let case_RCanR := tt  in RCanR _
-          | RuCanL a             => let case_RuCanL := tt in RuCanL _
-          | RuCanR a             => let case_RuCanR := tt in RuCanR _
-          | RAssoc a b c         => let case_RAssoc := tt in RAssoc _ _ _
-          | RCossa a b c         => let case_RCossa := tt in RCossa _ _ _
-          | RExch  a b           => let case_RExch := tt  in RExch _ _
-          | RWeak  a             => let case_RWeak := tt  in RWeak _
-          | RCont  a             => let case_RCont := tt  in RCont _
-          | RLeft  a b c r'      => let case_RLeft := tt  in RLeft  _ (flatten _ _ r')
-          | RRight a b c r'      => let case_RRight := tt in RRight _ (flatten _ _ r')
-          | RComp  a b c r1 r2   => let case_RComp := tt  in RComp    (flatten _ _ r1) (flatten _ _ r2)
+          | AId  a               => let case_AId := tt  in AId _
+          | ACanL  a             => let case_ACanL := tt  in ACanL _
+          | ACanR  a             => let case_ACanR := tt  in ACanR _
+          | AuCanL a             => let case_AuCanL := tt in AuCanL _
+          | AuCanR a             => let case_AuCanR := tt in AuCanR _
+          | AAssoc a b c         => let case_AAssoc := tt in AAssoc _ _ _
+          | AuAssoc a b c         => let case_AuAssoc := tt in AuAssoc _ _ _
+          | AExch  a b           => let case_AExch := tt  in AExch _ _
+          | AWeak  a             => let case_AWeak := tt  in AWeak _
+          | ACont  a             => let case_ACont := tt  in ACont _
+          | ALeft  a b c r'      => let case_ALeft := tt  in ALeft  _ (flatten _ _ r')
+          | ARight a b c r'      => let case_ARight := tt in ARight _ (flatten _ _ r')
+          | AComp  a b c r1 r2   => let case_AComp := tt  in AComp    (flatten _ _ r1) (flatten _ _ r2)
         end) ant1 ant2 r); clear flatten; repeat take_simplify; repeat drop_simplify; intros.
         Defined.
 
@@ -574,19 +511,19 @@ Section HaskFlattener.
       apply nd_rule.
       apply RArrange.
       induction r; simpl.
-        apply RId.
-        apply RCanL.
-        apply RCanR.
-        apply RuCanL.
-        apply RuCanR.
-        apply RAssoc.
-        apply RCossa.
-        apply RExch.    (* TO DO: check for all-leaf trees here *)
-        apply RWeak.
-        apply RCont.
-        apply RLeft; auto.
-        apply RRight; auto.
-        eapply RComp; [ apply IHr1 | apply IHr2 ].
+        apply AId.
+        apply ACanL.
+        apply ACanR.
+        apply AuCanL.
+        apply AuCanR.
+        apply AAssoc.
+        apply AuAssoc.
+        apply AExch.    (* TO DO: check for all-leaf trees here *)
+        apply AWeak.
+        apply ACont.
+        apply ALeft; auto.
+        apply ARight; auto.
+        eapply AComp; [ apply IHr1 | apply IHr2 ].
 
       apply flatten_arrangement.
         apply r.
@@ -609,11 +546,11 @@ Section HaskFlattener.
     eapply nd_comp; [ idtac | eapply nd_rule; eapply RLet ].
     eapply nd_comp; [ apply nd_llecnac | idtac ].
     apply nd_prod.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanL ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanL ].
     eapply nd_comp; [ idtac | eapply postcompose_ ].
     apply ga_uncancelr.
 
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     eapply nd_comp; [ idtac | eapply precompose ].
     apply pfb.
     Defined.
@@ -627,7 +564,7 @@ Section HaskFlattener.
 
     intros.
     unfold drop_lev.
-    set (@arrange' _ succ (levelMatch (ec::nil))) as q.
+    set (@arrangeUnPartition _ succ (levelMatch (ec::nil))) as q.
     set (arrangeMap _ _ flatten_leveled_type q) as y.
     eapply nd_comp.
     Focus 2.
@@ -636,7 +573,7 @@ Section HaskFlattener.
     apply y.
     idtac.
     clear y q.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
     simpl.
     eapply nd_comp; [ apply nd_llecnac | idtac ].
     eapply nd_comp; [ idtac | eapply nd_rule; eapply RLet ].
@@ -677,59 +614,6 @@ Section HaskFlattener.
       apply IHsucc2.
     Defined.
 
-  Definition arrange_empty_tree : forall {T}{A}(q:Tree A)(t:Tree ??T),
-    t = mapTree (fun _:A => None) q ->
-    Arrange t [].
-    intros T A q.
-    induction q; intros.
-      simpl in H.
-      rewrite H.
-      apply RId.
-    simpl in *.
-    destruct t; try destruct o; inversion H.
-      set (IHq1 _ H1) as x1.
-      set (IHq2 _ H2) as x2.
-      eapply RComp.
-        eapply RRight.
-        rewrite <- H1.
-        apply x1.
-      eapply RComp.
-        apply RCanL.
-        rewrite <- H2.
-        apply x2.
-      Defined.
-
-(*  Definition unarrange_empty_tree : forall {T}{A}(t:Tree ??T)(q:Tree A),
-    t = mapTree (fun _:A => None) q ->
-    Arrange [] t.
-  Defined.*)
-
-  Definition decide_tree_empty : forall {T:Type}(t:Tree ??T),
-    sum { q:Tree unit & t = mapTree (fun _ => None) q } unit.
-    intro T.
-    refine (fix foo t :=
-      match t with
-        | T_Leaf x => _
-        | T_Branch b1 b2 => let b1' := foo b1 in let b2' := foo b2 in _
-      end).
-    intros.
-    destruct x.
-    right; apply tt.
-    left.
-      exists (T_Leaf tt).
-      auto.
-    destruct b1'.
-    destruct b2'.
-    destruct s.
-    destruct s0.
-    subst.
-    left.
-    exists (x,,x0).
-    reflexivity.
-    right; auto.
-    right; auto.
-    Defined.
-
   Definition arrange_esc : forall Γ Δ ec succ t,
    ND Rule
      [Γ > Δ > mapOptionTree (flatten_leveled_type ) succ |- [t]@nil]
@@ -737,7 +621,7 @@ Section HaskFlattener.
       [(@ga_mk _ (v2t ec) [] (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: nil) succ))) @@  nil],,
       mapOptionTree (flatten_leveled_type ) (drop_lev (ec :: nil) succ)  |- [t]@nil].
     intros.
-    set (@arrange _ succ (levelMatch (ec::nil))) as q.
+    set (@arrangePartition _ succ (levelMatch (ec::nil))) as q.
     set (@drop_lev Γ (ec::nil) succ) as q'.
     assert (@drop_lev Γ (ec::nil) succ=q') as H.
       reflexivity.
@@ -757,7 +641,7 @@ Section HaskFlattener.
     destruct s.
 
     simpl.
-    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RExch ].
+    eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply AExch ].
     set (fun z z' => @RLet Γ Δ z (mapOptionTree flatten_leveled_type q') t z' nil) as q''.
     eapply nd_comp; [ idtac | eapply nd_rule; apply RLet ].
     clear q''.
@@ -765,19 +649,19 @@ Section HaskFlattener.
     apply nd_prod.
     apply nd_rule.
     apply RArrange.
-    eapply RComp; [ idtac | apply RCanR ].
-    apply RLeft.
-    apply (@arrange_empty_tree _ _ _ _ e).
+    eapply AComp; [ idtac | apply ACanR ].
+    apply ALeft.
+    apply (@arrangeCancelEmptyTree _ _ _ _ e).
    
     eapply nd_comp.
       eapply nd_rule.
       eapply (@RVar Γ Δ t nil).
     apply nd_rule.
       apply RArrange.
-      eapply RComp.
-      apply RuCanR.
-      apply RLeft.
-      apply RWeak.
+      eapply AComp.
+      apply AuCanR.
+      apply ALeft.
+      apply AWeak.
 (*
     eapply decide_tree_empty.
 
@@ -801,25 +685,19 @@ Section HaskFlattener.
         simpl.
         apply nd_rule.
         apply RArrange.
-        apply RLeft.
-        apply RWeak.
+        apply ALeft.
+        apply AWeak.
       simpl.
         apply nd_rule.
         unfold take_lev.
         simpl.
         apply RArrange.
-        apply RLeft.
-        apply RWeak.
+        apply ALeft.
+        apply AWeak.
       apply (Prelude_error "escapifying code with multi-leaf antecedents is not supported").
 *)
       Defined.
 
-  Lemma mapOptionTree_distributes
-    : forall T R (a b:Tree ??T) (f:T->R),
-      mapOptionTree f (a,,b) = (mapOptionTree f a),,(mapOptionTree f b).
-    reflexivity.
-    Qed.
-
   Lemma unlev_relev : forall {Γ}(t:Tree ??(HaskType Γ ★)) lev, mapOptionTree unlev (t @@@ lev) = t.
     intros.
     induction t.
@@ -836,17 +714,17 @@ Section HaskFlattener.
     simpl.
     drop_simplify.
     simpl.
-    apply RId.
+    apply AId.
     simpl.
-    apply RId.
-    eapply RComp; [ idtac | apply RCanL ].
-    eapply RComp; [ idtac | eapply RLeft; apply IHt2 ].
+    apply AId.
+    eapply AComp; [ idtac | apply ACanL ].
+    eapply AComp; [ idtac | eapply ALeft; apply IHt2 ].
     Opaque drop_lev.
     simpl.
     Transparent drop_lev.
     idtac.
     drop_simplify.
-    apply RRight.
+    apply ARight.
     apply IHt1.
     Defined.
 
@@ -857,17 +735,17 @@ Section HaskFlattener.
     simpl.
     drop_simplify.
     simpl.
-    apply RId.
+    apply AId.
     simpl.
-    apply RId.
-    eapply RComp; [ apply RuCanL | idtac ].
-    eapply RComp; [ eapply RRight; apply IHt1 | idtac ].
+    apply AId.
+    eapply AComp; [ apply AuCanL | idtac ].
+    eapply AComp; [ eapply ARight; apply IHt1 | idtac ].
     Opaque drop_lev.
     simpl.
     Transparent drop_lev.
     idtac.
     drop_simplify.
-    apply RLeft.
+    apply ALeft.
     apply IHt2.
     Defined.
 
@@ -906,7 +784,22 @@ Section HaskFlattener.
     admit.
     Qed.
 
-  Definition flatten_proof :
+  Lemma drop_to_nothing : forall (Γ:TypeEnv) Σ (lev:HaskLevel Γ),
+    drop_lev lev (Σ @@@ lev) = mapTree (fun _ => None) (mapTree (fun _ => tt) Σ).
+    intros.
+    induction Σ.
+    destruct a; simpl.
+    drop_simplify.
+    auto.
+    drop_simplify.
+    auto.
+    simpl.
+    rewrite <- IHΣ1.
+    rewrite <- IHΣ2.
+    reflexivity.
+    Qed.
+
+  Definition flatten_skolemized_proof :
     forall  {h}{c},
       ND SRule h c ->
       ND  Rule (mapOptionTree (flatten_judgment ) h) (mapOptionTree (flatten_judgment ) c).
@@ -938,6 +831,9 @@ Section HaskFlattener.
       | RAbsCo   Γ Δ Σ κ σ  σ₁ σ₂  lev => let case_RAbsCo := tt        in _
       | RApp     Γ Δ Σ₁ Σ₂ tx te lev   => let case_RApp := tt          in _
       | RLet     Γ Δ Σ₁ Σ₂ σ₁ σ₂ lev   => let case_RLet := tt          in _
+      | RCut    Γ Δ Σ₁ Σ₁₂ Σ₂ Σ₃ l    => let case_RCut := tt          in _
+      | RLeft   Γ Δ Σ₁ Σ₂  Σ     l    => let case_RLeft := tt in _
+      | RRight  Γ Δ Σ₁ Σ₂  Σ     l    => let case_RRight := tt in _
       | RWhere   Γ Δ Σ₁ Σ₂ Σ₃ σ₁ σ₂ lev   => let case_RWhere := tt          in _
       | RJoin    Γ p lri m x q l       => let case_RJoin := tt in _
       | RVoid    _ _       l           => let case_RVoid := tt   in _
@@ -1023,7 +919,7 @@ Section HaskFlattener.
         eapply nd_rule.
         eapply RArrange.
         simpl.
-        apply RCanR.
+        apply ACanR.
       apply boost.
       simpl.
       apply ga_curry.
@@ -1089,8 +985,8 @@ Section HaskFlattener.
       eapply nd_comp; [ idtac | eapply nd_rule; eapply RLet ].
       apply nd_prod.
       apply nd_id.
-      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply RCanL | idtac ].
-      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RExch (* okay *)].
+      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply ACanL | idtac ].
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch (* okay *)].
       apply precompose.
 
     destruct case_RWhere.
@@ -1119,8 +1015,127 @@ Section HaskFlattener.
       eapply nd_comp; [ idtac | eapply precompose' ].
       apply nd_rule.
       apply RArrange.
-      apply RLeft.
-      apply RCanL.
+      apply ALeft.
+      apply ACanL.
+
+    destruct case_RCut.
+      simpl.
+      destruct l as [|ec lev]; simpl.
+        apply nd_rule.
+        replace (mapOptionTree flatten_leveled_type (Σ₁₂ @@@ nil)) with (mapOptionTree flatten_type Σ₁₂ @@@ nil).
+        apply RCut.
+        induction Σ₁₂; try destruct a; auto.
+        simpl.
+        rewrite <- IHΣ₁₂1.
+        rewrite <- IHΣ₁₂2.
+        reflexivity.
+      simpl.
+      repeat drop_simplify.
+      simpl.
+      repeat take_simplify.
+      simpl.
+      set (drop_lev (ec :: lev) (Σ₁₂ @@@ (ec :: lev))) as x1.
+      rewrite take_lemma'.
+      rewrite mapOptionTree_compose.
+      rewrite mapOptionTree_compose.
+      rewrite mapOptionTree_compose.
+      rewrite unlev_relev.
+      rewrite <- mapOptionTree_compose.
+      rewrite <- mapOptionTree_compose.
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RCut ]. 
+      apply nd_prod.
+      apply nd_id.
+      eapply nd_comp.
+      eapply nd_rule.
+      eapply RArrange.
+      eapply ARight.
+      unfold x1.
+      rewrite drop_to_nothing.
+      apply arrangeCancelEmptyTree with (q:=(mapTree (fun _ : ??(HaskType Γ ★) => tt) Σ₁₂)).
+      admit. (* OK *)
+      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply ACanL | idtac ].
+      set (mapOptionTree flatten_type Σ₁₂) as a.
+      set (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) Σ₁)) as b.
+      set (mapOptionTree flatten_leveled_type (drop_lev (ec :: lev) Σ₂)) as c.
+      set (mapOptionTree (flatten_type ○ unlev) (take_lev (ec :: lev) Σ₂)) as d.
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RCut ].
+      eapply nd_comp; [ apply nd_llecnac | idtac ].
+      apply nd_prod.
+      simpl.
+      eapply ga_first.
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AExch ].
+      simpl.
+      apply precompose.
+
+    destruct case_RLeft.
+      simpl.
+      destruct l as [|ec lev].
+      simpl.
+        replace (mapOptionTree flatten_leveled_type (Σ @@@ nil)) with (mapOptionTree flatten_type Σ @@@ nil).
+        apply nd_rule.
+        apply RLeft.
+        induction Σ; try destruct a; auto.
+        simpl.
+        rewrite <- IHΣ1.
+        rewrite <- IHΣ2.
+        reflexivity.
+      repeat drop_simplify.
+        rewrite drop_to_nothing.
+        simpl.
+        eapply nd_comp.
+        Focus 2.
+        eapply nd_rule.
+        eapply RArrange.
+        eapply ARight.
+        apply arrangeUnCancelEmptyTree with (q:=(mapTree (fun _ : ??(HaskType Γ ★) => tt) Σ)).
+        admit (* FIXME *).
+        idtac.
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AuCanL ].
+        apply boost.
+        take_simplify.
+        simpl.
+        replace (take_lev (ec :: lev) (Σ @@@ (ec :: lev))) with (Σ @@@ (ec::lev)).
+        rewrite mapOptionTree_compose.
+        rewrite mapOptionTree_compose.
+        rewrite unlev_relev.
+        apply ga_second.
+      rewrite take_lemma'.
+      reflexivity.
+        
+    destruct case_RRight.
+      simpl.
+      destruct l as [|ec lev].
+      simpl.
+        replace (mapOptionTree flatten_leveled_type (Σ @@@ nil)) with (mapOptionTree flatten_type Σ @@@ nil).
+        apply nd_rule.
+        apply RRight.
+        induction Σ; try destruct a; auto.
+        simpl.
+        rewrite <- IHΣ1.
+        rewrite <- IHΣ2.
+        reflexivity.
+      repeat drop_simplify.
+        rewrite drop_to_nothing.
+        simpl.
+        eapply nd_comp.
+        Focus 2.
+        eapply nd_rule.
+        eapply RArrange.
+        eapply ALeft.
+        apply arrangeUnCancelEmptyTree with (q:=(mapTree (fun _ : ??(HaskType Γ ★) => tt) Σ)).
+        admit (* FIXME *).
+        idtac.
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply AuCanR ].
+        apply boost.
+        take_simplify.
+        simpl.
+        replace (take_lev (ec :: lev) (Σ @@@ (ec :: lev))) with (Σ @@@ (ec::lev)).
+        rewrite mapOptionTree_compose.
+        rewrite mapOptionTree_compose.
+        rewrite unlev_relev.
+        apply ga_first.
+      rewrite take_lemma'.
+      reflexivity.
 
     destruct case_RVoid.
       simpl.
@@ -1222,15 +1237,15 @@ Section HaskFlattener.
       set (mapOptionTree (flatten_type ○ unlev)(take_lev (ec :: nil) succ)) as succ_guest.
       set (mapOptionTree flatten_type (take_arg_types_as_tree t)) as succ_args.
       unfold empty_tree.
-      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply RLeft; apply tree_of_nothing | idtac ].
-      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply RCanR | idtac ].
+      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply ALeft; apply tree_of_nothing | idtac ].
+      eapply nd_comp; [ eapply nd_rule; eapply RArrange; eapply ACanR | idtac ].
       refine (ga_unkappa Γ Δ (v2t ec) nil _ _ _ _ ;; _).
       eapply nd_comp; [ idtac | eapply arrange_brak ].
       unfold succ_host.
       unfold succ_guest.
       eapply nd_rule.
         eapply RArrange.
-        apply RExch.
+        apply AExch.
       apply (Prelude_error "found Brak at depth >0 indicating 3-level code; only two-level code is currently supported").
 
     destruct case_SEsc.
@@ -1244,8 +1259,8 @@ Section HaskFlattener.
       take_simplify.
       drop_simplify.
       simpl.
-      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RLeft; apply tree_of_nothing' ].
-      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanR ].
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ALeft; apply tree_of_nothing' ].
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanR ].
       simpl.
       rewrite take_lemma'.
       rewrite unlev_relev.
@@ -1261,15 +1276,15 @@ Section HaskFlattener.
       set (mapOptionTree flatten_leveled_type (drop_lev (ec :: nil) succ)) as succ_host.
       set (mapOptionTree flatten_type (take_arg_types_as_tree t)) as succ_args.
 
-      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RuCanR ].
-      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RuCanR ].
-      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply RCanL ].
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply AuCanR ].
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply AuCanR ].
+      eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; apply ACanL ].
       eapply nd_comp; [ idtac | eapply nd_rule; eapply RLet ].
       eapply nd_comp; [ apply nd_llecnac | idtac ].
       apply nd_prod; [ idtac | eapply boost ].
       induction x.
         apply ga_id.
-        eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply RCanL ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply RArrange; eapply ACanL ].
         simpl.
         apply ga_join.
           apply IHx1.
@@ -1300,6 +1315,13 @@ Section HaskFlattener.
       apply (Prelude_error "found Esc at depth >0 indicating 3-level code; only two-level code is currently supported").
       Defined.
 
+  Definition flatten_proof :
+    forall  {h}{c},
+      ND  Rule h c ->
+      ND  Rule h c.
+    apply (Prelude_error "sorry, non-skolemized flattening isn't implemented").
+    Defined.
+
   Definition skolemize_and_flatten_proof :
     forall  {h}{c},
       ND  Rule h c ->
@@ -1309,7 +1331,7 @@ Section HaskFlattener.
     intros.
     rewrite mapOptionTree_compose.
     rewrite mapOptionTree_compose.
-    apply flatten_proof.
+    apply flatten_skolemized_proof.
     apply skolemize_proof.
     apply X.
     Defined.
index 8aba304..30a0ae8 100644 (file)
@@ -73,10 +73,10 @@ Section HaskProgrammingLanguage.
       apply nd_id.
       eapply nd_rule.
       set (@org_fc) as ofc.
-      set (RArrange Γ Δ _ _ _ (RuCanL [l0])) as rule.
-      apply org_fc with (r:=RArrange _ _ _ _ _ (RuCanL [_])).
+      set (RArrange Γ Δ _ _ _ (AuCanL [l0])) as rule.
+      apply org_fc with (r:=RArrange _ _ _ _ _ (AuCanL [_])).
       auto.
-      eapply nd_comp; [ idtac | eapply nd_rule; apply org_fc with (r:=RArrange _ _ _ _ _ (RCanL _)) ].
+      eapply nd_comp; [ idtac | eapply nd_rule; apply org_fc with (r:=RArrange _ _ _ _ _ (ACanL _)) ].
       apply nd_rule.
       destruct l.
       destruct l0.
@@ -148,23 +148,23 @@ Section HaskProgrammingLanguage.
   ; cnd_expand_right := fun a b c => SystemFCa_right c a b }.
     (*
     intros; apply nd_rule. simpl.
-      apply (org_fc _ _ _ _ ((RArrange _ _ _ _ _ (RCossa _ _ _)))).
+      apply (org_fc _ _ _ _ ((RArrange _ _ _ _ _ (AuAssoc _ _ _)))).
       auto.
 
     intros; apply nd_rule. simpl.
-      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RAssoc _ _ _))); auto.
+      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (AAssoc _ _ _))); auto.
 
     intros; apply nd_rule. simpl.
-      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RCanL _))); auto.
+      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (ACanL _))); auto.
 
     intros; apply nd_rule. simpl.
-      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RCanR _))); auto.
+      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (ACanR _))); auto.
 
     intros; apply nd_rule. simpl.
-      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RuCanL _))); auto.
+      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (AuCanL _))); auto.
 
     intros; apply nd_rule. simpl.
-      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (RuCanR _))); auto.
+      apply (org_fc _ _ _ _ (RArrange _ _ _ _ _ (AuCanR _))); auto.
       *)
       admit.
       admit.
index d92dd6b..46009f8 100644 (file)
@@ -10,6 +10,7 @@ Generalizable All Variables.
 Require Import Preamble.
 Require Import General.
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 Require Import Coq.Strings.String.
 Require Import Coq.Lists.List.
 Require Import HaskKinds.
@@ -44,23 +45,6 @@ Record ProofCaseBranch {tc:TyCon}{Γ}{Δ}{lev}{branchtype : HaskType Γ ★}{ava
 }.
 Implicit Arguments ProofCaseBranch [ ].
 
-(* Figure 3, production $\vdash_E$, Uniform rules *)
-Inductive Arrange {T} : Tree ??T -> Tree ??T -> Type :=
-| RId     : forall a        ,                Arrange           a                  a
-| RCanL   : forall a        ,                Arrange  (    [],,a   )      (       a   )
-| RCanR   : forall a        ,                Arrange  (    a,,[]   )      (       a   )
-| RuCanL  : forall a        ,                Arrange  (       a    )      (  [],,a    )
-| RuCanR  : forall a        ,                Arrange  (       a    )      (  a,,[]    )
-| RAssoc  : forall a b c    ,                Arrange  (a,,(b,,c)   )      ((a,,b),,c  )
-| RCossa  : forall a b c    ,                Arrange  ((a,,b),,c   )      ( a,,(b,,c) )
-| RExch   : forall a b      ,                Arrange  (   (b,,a)   )      (  (a,,b)   )
-| RWeak   : forall a        ,                Arrange  (       []   )      (       a   )
-| RCont   : forall a        ,                Arrange  (  (a,,a)    )      (       a   )
-| RLeft   : forall {h}{c} x , Arrange h c -> Arrange  (    x,,h    )      (       x,,c)
-| RRight  : forall {h}{c} x , Arrange h c -> Arrange  (    h,,x    )      (       c,,x)
-| RComp   : forall {a}{b}{c}, Arrange a b -> Arrange b c -> Arrange a c
-.
-
 (* Figure 3, production $\vdash_E$, all rules *)
 Inductive Rule : Tree ??Judg -> Tree ??Judg -> Type :=
 
@@ -83,12 +67,16 @@ Inductive Rule : Tree ??Judg -> Tree ??Judg -> Type :=
 
 | RJoin  : ∀ Γ Δ Σ₁ Σ₂ τ₁ τ₂ l,   Rule ([Γ > Δ > Σ₁ |- τ₁ @l],,[Γ > Δ > Σ₂ |- τ₂ @l])         [Γ>Δ>  Σ₁,,Σ₂  |- τ₁,,τ₂       @l ]
 
-(* order is important here; we want to be able to skolemize without introducing new RExch'es *)
+(* order is important here; we want to be able to skolemize without introducing new AExch'es *)
 | RApp           : ∀ Γ Δ Σ₁ Σ₂ tx te l,  Rule ([Γ>Δ> Σ₁ |- [tx--->te]@l],,[Γ>Δ> Σ₂ |- [tx]@l])  [Γ>Δ> Σ₁,,Σ₂ |- [te]@l]
 
 | RLet           : ∀ Γ Δ Σ₁ Σ₂ σ₁ σ₂ l,  Rule ([Γ>Δ> Σ₁ |- [σ₁]@l],,[Γ>Δ> [σ₁@@l],,Σ₂ |- [σ₂]@l ])     [Γ>Δ> Σ₁,,Σ₂ |- [σ₂   ]@l]
 | RWhere         : ∀ Γ Δ Σ₁ Σ₂ Σ₃ σ₁ σ₂ l,  Rule ([Γ>Δ> Σ₁,,([σ₁@@l],,Σ₃) |- [σ₂]@l ],,[Γ>Δ> Σ₂ |- [σ₁]@l])     [Γ>Δ> Σ₁,,(Σ₂,,Σ₃) |- [σ₂ ]@l]
 
+| RCut           : ∀ Γ Δ Σ₁ Σ₁₂ Σ₂ Σ₃ l,  Rule ([Γ>Δ> Σ₁ |- Σ₁₂ @l],,[Γ>Δ> (Σ₁₂@@@l),,Σ₂ |- Σ₃@l ]) [Γ>Δ> Σ₁,,Σ₂ |- Σ₃@l]
+| RLeft          : ∀ Γ Δ Σ₁ Σ₂  Σ     l,  Rule  [Γ>Δ> Σ₁ |- Σ₂  @l]                                 [Γ>Δ> (Σ@@@l),,Σ₁ |- Σ,,Σ₂@l]
+| RRight         : ∀ Γ Δ Σ₁ Σ₂  Σ     l,  Rule  [Γ>Δ> Σ₁ |- Σ₂  @l]                                 [Γ>Δ> Σ₁,,(Σ@@@l) |- Σ₂,,Σ@l]
+
 | RVoid    : ∀ Γ Δ l,               Rule [] [Γ > Δ > [] |- [] @l ]
 
 | RAppT   : forall Γ Δ Σ κ σ (τ:HaskType Γ κ) l,      Rule [Γ>Δ> Σ   |- [HaskTAll κ σ]@l]      [Γ>Δ>    Σ     |- [substT σ τ]@l]
@@ -163,6 +151,9 @@ Lemma no_rules_with_multiple_conclusions : forall c h,
     destruct X0; destruct s; inversion e.
     destruct X0; destruct s; inversion e.
     destruct X0; destruct s; inversion e.
+    destruct X0; destruct s; inversion e.
+    destruct X0; destruct s; inversion e.
+    destruct X0; destruct s; inversion e.
     Qed.
 
 Lemma systemfc_all_rules_one_conclusion : forall h c1 c2 (r:Rule h (c1,,c2)), False.
@@ -174,41 +165,3 @@ Lemma systemfc_all_rules_one_conclusion : forall h c1 c2 (r:Rule h (c1,,c2)), Fa
   auto.
   Qed.
 
-(* "Arrange" objects are parametric in the type of the leaves of the tree *)
-Definition arrangeMap :
-  forall {T} (Σ₁ Σ₂:Tree ??T) {R} (f:T -> R),
-    Arrange Σ₁ Σ₂ ->
-    Arrange (mapOptionTree f Σ₁) (mapOptionTree f Σ₂).
-  intros.
-  induction X; simpl.
-  apply RId.
-  apply RCanL.
-  apply RCanR.
-  apply RuCanL.
-  apply RuCanR.
-  apply RAssoc.
-  apply RCossa.
-  apply RExch.
-  apply RWeak.
-  apply RCont.
-  apply RLeft; auto.
-  apply RRight; auto.
-  eapply RComp; [ apply IHX1 | apply IHX2 ].
-  Defined.
-
-(* a frequently-used Arrange *)
-Definition arrangeSwapMiddle {T} (a b c d:Tree ??T) :
-  Arrange ((a,,b),,(c,,d)) ((a,,c),,(b,,d)).
-  eapply RComp.
-  apply RCossa.
-  eapply RComp.
-  eapply RLeft.
-  eapply RComp.
-  eapply RAssoc.
-  eapply RRight.
-  apply RExch.
-  eapply RComp.
-  eapply RLeft.
-  eapply RCossa.
-  eapply RAssoc.
-  Defined.
index 4319a1c..dedc152 100644 (file)
@@ -6,6 +6,7 @@ Generalizable All Variables.
 Require Import Preamble.
 Require Import General.
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 Require Import Coq.Strings.String.
 Require Import Coq.Lists.List.
 Require Import HaskKinds.
@@ -159,19 +160,19 @@ Instance ToLatexMathJudgment : ToLatexMath Judg :=
 
 Fixpoint nd_uruleToRawLatexMath {T}{h}{c}(r:@Arrange T h c) : string :=
   match r with
-    | RLeft   _ _ _ r => nd_uruleToRawLatexMath r
-    | RRight  _ _ _ r => nd_uruleToRawLatexMath r
-    | RId     _     => "Id"
-    | RCanL   _     => "CanL"
-    | RCanR   _     => "CanR"
-    | RuCanL  _     => "uCanL"
-    | RuCanR  _     => "uCanR"
-    | RAssoc  _ _ _ => "Assoc"
-    | RCossa  _ _ _ => "Cossa"
-    | RExch   _ _   => "Exch"
-    | RWeak   _     => "Weak"
-    | RCont   _     => "Cont"
-    | RComp   _ _ _ _ _  => "Comp"  (* FIXME: do a better job here *)
+    | ALeft   _ _ _ r => nd_uruleToRawLatexMath r
+    | ARight  _ _ _ r => nd_uruleToRawLatexMath r
+    | AId     _     => "Id"
+    | ACanL   _     => "CanL"
+    | ACanR   _     => "CanR"
+    | AuCanL  _     => "uCanL"
+    | AuCanR  _     => "uCanR"
+    | AAssoc  _ _ _ => "Assoc"
+    | AuAssoc  _ _ _ => "Cossa"
+    | AExch   _ _   => "Exch"
+    | AWeak   _     => "Weak"
+    | ACont   _     => "Cont"
+    | AComp   _ _ _ _ _  => "Comp"  (* FIXME: do a better job here *)
   end.
 
 Fixpoint nd_ruleToRawLatexMath {h}{c}(r:Rule h c) : string :=
@@ -189,6 +190,9 @@ Fixpoint nd_ruleToRawLatexMath {h}{c}(r:Rule h c) : string :=
     | RAbsCo        _ _ _ _ _ _ _ _   => "AbsCo"
     | RApp          _ _ _ _ _ _ _     => "App"
     | RLet          _ _ _ _ _ _ _     => "Let"
+    | RCut          _ _ _ _ _ _ _     => "Cut"
+    | RLeft         _ _ _ _ _ _       => "Left"
+    | RRight        _ _ _ _ _ _       => "Right"
     | RWhere        _ _ _ _ _ _ _ _   => "Where"
     | RJoin         _ _ _ _ _ _ _     => "RJoin"
     | RLetRec       _ _ _ _ _ _       => "LetRec"
@@ -200,19 +204,19 @@ end.
 
 Fixpoint nd_hideURule {T}{h}{c}(r:@Arrange T h c) : bool :=
   match r with
-    | RLeft   _ _ _ r             => nd_hideURule r
-    | RRight  _ _ _ r             => nd_hideURule r
-    | RCanL   _                   => true
-    | RCanR   _                   => true
-    | RuCanL  _                   => true
-    | RuCanR  _                   => true
-    | RAssoc  _ _ _               => true
-    | RCossa  _ _ _               => true
-    | RExch      (T_Leaf None) b  => true
-    | RExch   a  (T_Leaf None)    => true
-    | RWeak      (T_Leaf None)    => true
-    | RCont      (T_Leaf None)    => true
-    | RComp   _ _ _ _ _           => false   (* FIXME: do better *)
+    | ALeft   _ _ _ r             => nd_hideURule r
+    | ARight  _ _ _ r             => nd_hideURule r
+    | ACanL   _                   => true
+    | ACanR   _                   => true
+    | AuCanL  _                   => true
+    | AuCanR  _                   => true
+    | AAssoc  _ _ _               => true
+    | AuAssoc  _ _ _               => true
+    | AExch      (T_Leaf None) b  => true
+    | AExch   a  (T_Leaf None)    => true
+    | AWeak      (T_Leaf None)    => true
+    | ACont      (T_Leaf None)    => true
+    | AComp   _ _ _ _ _           => false   (* FIXME: do better *)
     | _                           => false
   end.
 Fixpoint nd_hideRule {h}{c}(r:Rule h c) : bool :=
index 78c2e41..b16d4a8 100644 (file)
@@ -6,6 +6,7 @@ Generalizable All Variables.
 Require Import Preamble.
 Require Import General.
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 Require Import Coq.Strings.String.
 Require Import Coq.Lists.List.
 Require Import Coq.Init.Specif.
@@ -28,10 +29,10 @@ Section HaskProofToStrong.
   Definition judg2exprType (j:Judg) : Type :=
     match j with
       (Γ > Δ > Σ |- τ @ l) => forall (ξ:ExprVarResolver Γ) vars, Σ = mapOptionTree ξ vars ->
-        FreshM (ITree _ (fun t => Expr Γ Δ ξ (t @@ l)) τ)
+        FreshM (ITree _ (fun t => Expr Γ Δ ξ t l) τ)
       end.
 
-  Definition justOne Γ Δ ξ τ : ITree _ (fun t => Expr Γ Δ ξ t) [τ] -> Expr Γ Δ ξ τ.
+  Definition justOne Γ Δ ξ τ l : ITree _ (fun t => Expr Γ Δ ξ t l) [τ] -> Expr Γ Δ ξ τ l.
     intros.
     inversion X; auto.
     Defined.
@@ -223,7 +224,7 @@ Section HaskProofToStrong.
     Defined.
 
   Definition ujudg2exprType Γ (ξ:ExprVarResolver Γ)(Δ:CoercionEnv Γ) Σ τ l : Type :=
-    forall vars, Σ = mapOptionTree ξ vars -> FreshM (ITree _ (fun t => Expr Γ Δ ξ (t@@l)) τ).
+    forall vars, Σ = mapOptionTree ξ vars -> FreshM (ITree _ (fun t => Expr Γ Δ ξ t l) τ).
 
   Definition urule2expr  : forall Γ Δ h j t l (r:@Arrange _ h j) (ξ:VV -> LeveledHaskType Γ ★),
     ujudg2exprType Γ ξ Δ h t l ->
@@ -237,49 +238,49 @@ Section HaskProofToStrong.
     ujudg2exprType Γ ξ Δ H t l ->
     ujudg2exprType Γ ξ Δ C t l
  with
-          | RLeft   h c ctx r => let case_RLeft  := tt in (fun e => _) (urule2expr _ _ _ _ r)
-          | RRight  h c ctx r => let case_RRight := tt in (fun e => _) (urule2expr _ _ _ _ r)
-          | RId     a       => let case_RId    := tt in _
-          | RCanL   a       => let case_RCanL  := tt in _
-          | RCanR   a       => let case_RCanR  := tt in _
-          | RuCanL  a       => let case_RuCanL := tt in _
-          | RuCanR  a       => let case_RuCanR := tt in _
-          | RAssoc  a b c   => let case_RAssoc := tt in _
-          | RCossa  a b c   => let case_RCossa := tt in _
-          | RExch   a b     => let case_RExch  := tt in _
-          | RWeak   a       => let case_RWeak  := tt in _
-          | RCont   a       => let case_RCont  := tt in _
-          | RComp   a b c f g => let case_RComp  := tt in (fun e1 e2 => _) (urule2expr _ _ _ _ f) (urule2expr _ _ _ _ g)
+          | ALeft   h c ctx r => let case_ALeft  := tt in (fun e => _) (urule2expr _ _ _ _ r)
+          | ARight  h c ctx r => let case_ARight := tt in (fun e => _) (urule2expr _ _ _ _ r)
+          | AId     a       => let case_AId    := tt in _
+          | ACanL   a       => let case_ACanL  := tt in _
+          | ACanR   a       => let case_ACanR  := tt in _
+          | AuCanL  a       => let case_AuCanL := tt in _
+          | AuCanR  a       => let case_AuCanR := tt in _
+          | AAssoc  a b c   => let case_AAssoc := tt in _
+          | AuAssoc  a b c   => let case_AuAssoc := tt in _
+          | AExch   a b     => let case_AExch  := tt in _
+          | AWeak   a       => let case_AWeak  := tt in _
+          | ACont   a       => let case_ACont  := tt in _
+          | AComp   a b c f g => let case_AComp  := tt in (fun e1 e2 => _) (urule2expr _ _ _ _ f) (urule2expr _ _ _ _ g)
           end); clear urule2expr; intros.
 
-      destruct case_RId.
+      destruct case_AId.
         apply X.
 
-      destruct case_RCanL.
+      destruct case_ACanL.
         simpl; unfold ujudg2exprType; intros.
         simpl in X.
         apply (X ([],,vars)).
         simpl; rewrite <- H; auto.
 
-      destruct case_RCanR.
+      destruct case_ACanR.
         simpl; unfold ujudg2exprType; intros.
         simpl in X.
         apply (X (vars,,[])).
         simpl; rewrite <- H; auto.
 
-      destruct case_RuCanL.
+      destruct case_AuCanL.
         simpl; unfold ujudg2exprType; intros.
         destruct vars; try destruct o; inversion H.
         simpl in X.
         apply (X vars2); auto.
 
-      destruct case_RuCanR.
+      destruct case_AuCanR.
         simpl; unfold ujudg2exprType; intros.
         destruct vars; try destruct o; inversion H.
         simpl in X.
         apply (X vars1); auto.
 
-      destruct case_RAssoc.
+      destruct case_AAssoc.
         simpl; unfold ujudg2exprType; intros.
         simpl in X.
         destruct vars; try destruct o; inversion H.
@@ -287,7 +288,7 @@ Section HaskProofToStrong.
         apply (X (vars1_1,,(vars1_2,,vars2))).
         subst; auto.
 
-      destruct case_RCossa.
+      destruct case_AuAssoc.
         simpl; unfold ujudg2exprType; intros.
         simpl in X.
         destruct vars; try destruct o; inversion H.
@@ -295,20 +296,20 @@ Section HaskProofToStrong.
         apply (X ((vars1,,vars2_1),,vars2_2)).
         subst; auto.
 
-      destruct case_RExch.
+      destruct case_AExch.
         simpl; unfold ujudg2exprType ; intros.
         simpl in X.
         destruct vars; try destruct o; inversion H.
         apply (X (vars2,,vars1)).
         inversion H; subst; auto.
         
-      destruct case_RWeak.
+      destruct case_AWeak.
         simpl; unfold ujudg2exprType; intros.
         simpl in X.
         apply (X []).
         auto.
         
-      destruct case_RCont.
+      destruct case_ACont.
         simpl; unfold ujudg2exprType ; intros.
         simpl in X.
         apply (X (vars,,vars)).
@@ -316,7 +317,7 @@ Section HaskProofToStrong.
         rewrite <- H.
         auto.
 
-      destruct case_RLeft.
+      destruct case_ALeft.
         intro vars; unfold ujudg2exprType; intro H.
         destruct vars; try destruct o; inversion H.
         apply (fun q => e ξ q vars2 H2).
@@ -331,7 +332,7 @@ Section HaskProofToStrong.
           simpl.
           reflexivity.
 
-      destruct case_RRight.
+      destruct case_ARight.
         intro vars; unfold ujudg2exprType; intro H.
         destruct vars; try destruct o; inversion H.
         apply (fun q => e ξ q vars1 H1).
@@ -346,7 +347,7 @@ Section HaskProofToStrong.
           simpl.
           reflexivity.
 
-      destruct case_RComp.
+      destruct case_AComp.
         apply e2.
         apply e1.
         apply X.
@@ -354,7 +355,7 @@ Section HaskProofToStrong.
 
   Definition letrec_helper Γ Δ l (varstypes:Tree ??(VV * HaskType Γ ★)) ξ' :
     ITree (HaskType Γ ★)
-         (fun t : HaskType Γ ★ => Expr Γ Δ ξ' (t @@ l))
+         (fun t : HaskType Γ ★ => Expr Γ Δ ξ' t l)
          (mapOptionTree (unlev ○ ξ' ○ (@fst _ _)) varstypes)
          -> ELetRecBindings Γ Δ ξ' l varstypes.
     intros.
@@ -421,7 +422,8 @@ Section HaskProofToStrong.
      prod (judg2exprType (pcb_judg (projT2 pcb))) {vars' : Tree ??VV & pcb_freevars (projT2 pcb) = mapOptionTree ξ vars'} ->
      ((fun sac => FreshM
        { scb : StrongCaseBranchWithVVs VV eqdec_vv tc avars sac
-         & Expr (sac_gamma sac Γ) (sac_delta sac Γ avars (weakCK'' Δ)) (scbwv_xi scb ξ lev) (weakLT' (tbranches @@  lev)) }) (projT1 pcb)).
+         & Expr (sac_gamma sac Γ) (sac_delta sac Γ avars (weakCK'' Δ)) (scbwv_xi scb ξ lev)
+         (weakT' tbranches) (weakL' lev) }) (projT1 pcb)).
     intro pcb.
     intro X.
     simpl in X.
@@ -507,6 +509,242 @@ Section HaskProofToStrong.
 
     Defined.
 
+  Lemma manyFresh : forall Γ Σ (ξ0:VV -> LeveledHaskType Γ ★),
+    FreshM { vars : _ & { ξ : VV -> LeveledHaskType Γ ★ & Σ = mapOptionTree ξ vars } }.
+    intros Γ Σ.
+    induction Σ; intro ξ.
+    destruct a.
+    destruct l as [τ l].
+    set (fresh_lemma' Γ [τ] [] [] ξ l (refl_equal _)) as q.
+    refine (q >>>= fun q' => return _).
+    apply FreshMon.
+    clear q.
+    destruct q' as [varstypes [pf1 [pf2 distpf]]].
+    exists (mapOptionTree (@fst _ _) varstypes).
+    exists (update_xi ξ l (leaves varstypes)).
+    symmetry; auto.
+    refine (return _).
+    exists [].
+    exists ξ; auto.
+    refine (bind f1 = IHΣ1 ξ ; _).
+    apply FreshMon.
+    destruct f1 as [vars1 [ξ1 pf1]].
+    refine (bind f2 = IHΣ2 ξ1 ; _).
+    apply FreshMon.
+    destruct f2 as [vars2 [ξ2 pf22]].
+    refine (return _).
+    exists (vars1,,vars2).
+    exists ξ2.
+    simpl.
+    rewrite pf22.
+    rewrite pf1.
+    admit.         (* freshness assumption *)
+    Defined.
+
+  Definition rlet Γ Δ Σ₁ Σ₂ σ₁ σ₂ p :
+    forall (X_ : ITree Judg judg2exprType
+         ([Γ > Δ > Σ₁ |- [σ₁] @ p],, [Γ > Δ > [σ₁ @@  p],, Σ₂ |- [σ₂] @ p])),
+   ITree Judg judg2exprType [Γ > Δ > Σ₁,, Σ₂ |- [σ₂] @ p].
+    intros.
+    apply ILeaf.
+    simpl in *; intros.
+    destruct vars; try destruct o; inversion H.
+
+    refine (fresh_lemma _ ξ _ _ σ₁ p H2 >>>= (fun pf => _)).
+    apply FreshMon.
+
+    destruct pf as [ vnew [ pf1 pf2 ]].
+    set (update_xi ξ p (((vnew, σ₁ )) :: nil)) as ξ' in *.
+    inversion X_.
+    apply ileaf in X.
+    apply ileaf in X0.
+    simpl in *.
+
+    refine (X ξ vars1 _ >>>= fun X0' => _).
+    apply FreshMon.
+    simpl.
+    auto.
+
+    refine (X0 ξ' ([vnew],,vars2) _ >>>= fun X1' => _).
+    apply FreshMon.
+    simpl.
+    rewrite pf2.
+    rewrite pf1.
+    reflexivity.
+    apply FreshMon.
+
+    apply ILeaf.
+    apply ileaf in X1'.
+    apply ileaf in X0'.
+    simpl in *.
+    apply ELet with (ev:=vnew)(tv:=σ₁).
+    apply X0'.
+    apply X1'.
+    Defined.
+
+  Definition vartree Γ Δ Σ lev ξ :
+    forall vars, Σ @@@ lev = mapOptionTree ξ vars ->
+    ITree (HaskType Γ ★) (fun t : HaskType Γ ★ => Expr Γ Δ ξ t lev) Σ.
+    induction Σ; intros.
+    destruct a.
+    intros; simpl in *.
+    apply ILeaf.
+    destruct vars; try destruct o; inversion H.
+    set (EVar Γ Δ ξ v) as q.
+    rewrite <- H1 in q.
+    apply q.
+    intros.
+    apply INone.
+    intros.
+    destruct vars; try destruct o; inversion H.
+    apply IBranch.
+    eapply IHΣ1.
+    apply H1.
+    eapply IHΣ2.
+    apply H2.
+    Defined.
+
+
+  Definition rdrop  Γ Δ Σ₁ Σ₁₂ a lev :
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |- a,,Σ₁₂ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |- a @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *.
+    intros.
+    set (X ξ vars H) as q.
+    simpl in q.
+    refine (q >>>= fun q' => return _).
+    apply FreshMon.
+    inversion q'.
+    apply X0.
+    Defined.
+
+  Definition rdrop'  Γ Δ Σ₁ Σ₁₂ a lev :
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |- Σ₁₂,,a @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |- a @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *.
+    intros.
+    set (X ξ vars H) as q.
+    simpl in q.
+    refine (q >>>= fun q' => return _).
+    apply FreshMon.
+    inversion q'.
+    auto.
+    Defined.
+
+  Definition rdrop''  Γ Δ Σ₁ Σ₁₂ lev :
+    ITree Judg judg2exprType [Γ > Δ > [],,Σ₁ |- Σ₁₂ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |- Σ₁₂ @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *; intros.
+    eapply X with (vars:=[],,vars).
+    rewrite H; reflexivity.
+    Defined.
+
+  Definition rdrop'''  Γ Δ a Σ₁ Σ₁₂ lev :
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |- Σ₁₂ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > a,,Σ₁ |- Σ₁₂ @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *; intros.
+    destruct vars; try destruct o; inversion H.
+    eapply X with (vars:=vars2).
+    auto.
+    Defined.
+
+  Definition rassoc  Γ Δ Σ₁ a b c lev :
+    ITree Judg judg2exprType [Γ > Δ > ((a,,b),,c) |- Σ₁ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > (a,,(b,,c)) |- Σ₁ @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *; intros.
+    destruct vars; try destruct o; inversion H.
+    destruct vars2; try destruct o; inversion H2.
+    apply X with (vars:=(vars1,,vars2_1),,vars2_2).
+    subst; reflexivity.
+    Defined.
+
+  Definition rassoc'  Γ Δ Σ₁ a b c lev :
+    ITree Judg judg2exprType [Γ > Δ > (a,,(b,,c)) |- Σ₁ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > ((a,,b),,c) |- Σ₁ @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *; intros.
+    destruct vars; try destruct o; inversion H.
+    destruct vars1; try destruct o; inversion H1.
+    apply X with (vars:=vars1_1,,(vars1_2,,vars2)).
+    subst; reflexivity.
+    Defined.
+
+  Definition swapr  Γ Δ Σ₁ a b c lev :
+    ITree Judg judg2exprType [Γ > Δ > ((a,,b),,c) |- Σ₁ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > ((b,,a),,c) |- Σ₁ @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *; intros.
+    destruct vars; try destruct o; inversion H.
+    destruct vars1; try destruct o; inversion H1.
+    apply X with (vars:=(vars1_2,,vars1_1),,vars2).
+    subst; reflexivity.
+    Defined.
+
+  Definition rdup  Γ Δ Σ₁ a  c lev :
+    ITree Judg judg2exprType [Γ > Δ > ((a,,a),,c) |- Σ₁ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > (a,,c) |- Σ₁ @ lev].
+    intros.
+    apply ileaf in X.
+    apply ILeaf.
+    simpl in *; intros.
+    destruct vars; try destruct o; inversion H.
+    apply X with (vars:=(vars1,,vars1),,vars2).    (* is this allowed? *)
+    subst; reflexivity.
+    Defined.
+
+  (* holy cow this is ugly *)
+  Definition rcut Γ Δ  Σ₃ lev  Σ₁₂  :
+    forall Σ₁ Σ₂,
+    ITree Judg judg2exprType [Γ > Δ > Σ₁ |-  Σ₁₂ @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ >  Σ₁₂ @@@ lev,,Σ₂ |- [Σ₃] @ lev] ->
+    ITree Judg judg2exprType [Γ > Δ > Σ₁,,Σ₂ |- [Σ₃] @ lev].
+
+    induction Σ₁₂.
+    intros.
+    destruct a.
+
+    eapply rlet.
+    apply IBranch.
+    apply X.
+    apply X0.
+
+    simpl in X0.
+    apply rdrop'' in X0.
+    apply rdrop'''.
+    apply X0.
+
+    intros.
+    simpl in X0.
+    apply rassoc in X0.
+    set (IHΣ₁₂1 _ _ (rdrop  _ _ _ _ _ _ X) X0) as q.
+    set (IHΣ₁₂2 _ (Σ₁,,Σ₂) (rdrop' _ _ _ _ _ _ X)) as q'.
+    apply rassoc' in q.
+    apply swapr in q.
+    apply rassoc in q.
+    set (q' q) as q''.
+    apply rassoc' in q''.
+    apply rdup in q''.
+    apply q''.
+    Defined.
 
   Definition rule2expr : forall h j (r:Rule h j), ITree _ judg2exprType h -> ITree _ judg2exprType j.
 
@@ -526,6 +764,9 @@ Section HaskProofToStrong.
       | RAbsCo  Γ Δ Σ κ σ  σ₁ σ₂  y   => let case_RAbsCo := tt        in _
       | RApp    Γ Δ Σ₁ Σ₂ tx te p     => let case_RApp := tt          in _
       | RLet    Γ Δ Σ₁ Σ₂ σ₁ σ₂ p     => let case_RLet := tt          in _
+      | RCut    Γ Δ Σ₁ Σ₁₂ Σ₂ Σ₃ l    => let case_RCut := tt          in _
+      | RLeft   Γ Δ Σ₁ Σ₂  Σ     l    => let case_RLeft := tt in _
+      | RRight  Γ Δ Σ₁ Σ₂  Σ     l    => let case_RRight := tt in _
       | RWhere  Γ Δ Σ₁ Σ₂ Σ₃ σ₁ σ₂ p  => let case_RWhere := tt          in _
       | RJoin   Γ p lri m x q l       => let case_RJoin := tt in _
       | RVoid   _ _ l                 => let case_RVoid := tt   in _
@@ -567,8 +808,11 @@ Section HaskProofToStrong.
 
   destruct case_RVar.
     apply ILeaf; simpl; intros; refine (return ILeaf _ _).
-    destruct vars; simpl in H; inversion H; destruct o. inversion H1. rewrite H2.
-    apply EVar.
+    destruct vars; simpl in H; inversion H; destruct o. inversion H1.
+    set (@EVar _ _ _ Δ ξ v) as q.
+    rewrite <- H2 in q.
+    simpl in q.
+    apply q.
     inversion H.
 
   destruct case_RGlobal.
@@ -640,40 +884,8 @@ Section HaskProofToStrong.
     apply (EApp _ _ _ _ _ _ q1' q2').
 
   destruct case_RLet.
-    apply ILeaf.
-    simpl in *; intros.
-    destruct vars; try destruct o; inversion H.
-
-    refine (fresh_lemma _ ξ _ _ σ₁ p H2 >>>= (fun pf => _)).
-    apply FreshMon.
-
-    destruct pf as [ vnew [ pf1 pf2 ]].
-    set (update_xi ξ p (((vnew, σ₁ )) :: nil)) as ξ' in *.
-    inversion X_.
-    apply ileaf in X.
-    apply ileaf in X0.
-    simpl in *.
-
-    refine (X ξ vars1 _ >>>= fun X0' => _).
-    apply FreshMon.
-    simpl.
-    auto.
-
-    refine (X0 ξ' ([vnew],,vars2) _ >>>= fun X1' => _).
-    apply FreshMon.
-    simpl.
-    rewrite pf2.
-    rewrite pf1.
-    reflexivity.
-    apply FreshMon.
-
-    apply ILeaf.
-    apply ileaf in X1'.
-    apply ileaf in X0'.
-    simpl in *.
-    apply ELet with (ev:=vnew)(tv:=σ₁).
-    apply X0'.
-    apply X1'.
+    eapply rlet.
+    apply X_.
 
   destruct case_RWhere.
     apply ILeaf.
@@ -715,6 +927,57 @@ Section HaskProofToStrong.
     apply X1'.
     apply X0'.
 
+  destruct case_RCut.
+    inversion X_.
+    subst.
+    clear X_.
+    induction Σ₃.
+    destruct a.
+    subst.
+    eapply rcut.
+    apply X.
+    apply X0.
+
+    apply ILeaf.
+    simpl.
+    intros.
+    refine (return _).
+    apply INone.
+    set (IHΣ₃1 (rdrop  _ _ _ _ _ _ X0)) as q1.
+    set (IHΣ₃2 (rdrop' _ _ _ _ _ _ X0)) as q2.
+    apply ileaf in q1.
+    apply ileaf in q2.
+    simpl in *.
+    apply ILeaf.
+    simpl.
+    intros.
+    refine (q1 _ _ H >>>= fun q1' => q2 _ _ H >>>= fun q2' => return _).
+    apply FreshMon.
+    apply FreshMon.
+    apply IBranch; auto.
+
+  destruct case_RLeft.
+    apply ILeaf.
+    simpl; intros.
+    destruct vars; try destruct o; inversion H.
+    refine (X_ _ _ H2 >>>= fun X' => return _).
+    apply FreshMon.
+    apply IBranch.
+    eapply vartree.
+    apply H1.
+    apply X'.
+
+  destruct case_RRight.
+    apply ILeaf.
+    simpl; intros.
+    destruct vars; try destruct o; inversion H.
+    refine (X_ _ _ H1 >>>= fun X' => return _).
+    apply FreshMon.
+    apply IBranch.
+    apply X'.
+    eapply vartree.
+    apply H2.
+
   destruct case_RVoid.
     apply ILeaf; simpl; intros.
     refine (return _).
@@ -820,41 +1083,9 @@ Section HaskProofToStrong.
     | scnd_branch _ _ _ c1 c2   => let case_branch := tt in fun q => IBranch _ _ (closed2expr _ _ c1 q) (closed2expr _ _ c2 q)
     end.
 
-  Lemma manyFresh : forall Γ Σ (ξ0:VV -> LeveledHaskType Γ ★),
-    FreshM { vars : _ & { ξ : VV -> LeveledHaskType Γ ★ & Σ = mapOptionTree ξ vars } }.
-    intros Γ Σ.
-    induction Σ; intro ξ.
-    destruct a.
-    destruct l as [τ l].
-    set (fresh_lemma' Γ [τ] [] [] ξ l (refl_equal _)) as q.
-    refine (q >>>= fun q' => return _).
-    apply FreshMon.
-    clear q.
-    destruct q' as [varstypes [pf1 [pf2 distpf]]].
-    exists (mapOptionTree (@fst _ _) varstypes).
-    exists (update_xi ξ l (leaves varstypes)).
-    symmetry; auto.
-    refine (return _).
-    exists [].
-    exists ξ; auto.
-    refine (bind f1 = IHΣ1 ξ ; _).
-    apply FreshMon.
-    destruct f1 as [vars1 [ξ1 pf1]].
-    refine (bind f2 = IHΣ2 ξ1 ; _).
-    apply FreshMon.
-    destruct f2 as [vars2 [ξ2 pf22]].
-    refine (return _).
-    exists (vars1,,vars2).
-    exists ξ2.
-    simpl.
-    rewrite pf22.
-    rewrite pf1.
-    admit.
-    Defined.
-
-  Definition proof2expr Γ Δ τ Σ (ξ0: VV -> LeveledHaskType Γ ★)
-    {zz:ToString VV} : ND Rule [] [Γ > Δ > Σ |- [unlev τ] @ getlev τ] ->
-    FreshM (???{ ξ : _ & Expr Γ Δ ξ τ}).
+  Definition proof2expr Γ Δ τ l Σ (ξ0: VV -> LeveledHaskType Γ ★)
+    {zz:ToString VV} : ND Rule [] [Γ > Δ > Σ |- [τ] @ l] ->
+    FreshM (???{ ξ : _ & Expr Γ Δ ξ τ l}).
     intro pf.
     set (mkSIND systemfc_all_rules_one_conclusion _ _ _ pf (scnd_weak [])) as cnd.
     apply closed2expr in cnd.
@@ -872,7 +1103,6 @@ Section HaskProofToStrong.
     exists ξ.
     apply ileaf in it.
     simpl in it.
-    destruct τ.
     apply it.
     apply INone.
     Defined.
index bb4dc92..6b6c756 100644 (file)
@@ -9,6 +9,7 @@ Generalizable All Variables.
 Require Import Preamble.
 Require Import General.
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 Require Import Coq.Strings.String.
 Require Import Coq.Lists.List.
 
@@ -124,14 +125,14 @@ Section HaskSkolemizer.
     destruct (eqd_dec ([tx],,take_arg_types_as_tree te) (take_arg_types_as_tree (tx ---> te))).
       rewrite <- e.
       simpl.
-      apply RId.
+      apply AId.
     unfold take_arg_types_as_tree.
       Opaque take_arg_types_as_tree.
       simpl.
       destruct (count_arg_types (te (fun _ : Kind => unit) (ite_unit Γ))).
       simpl.
       replace (tx) with (fun (TV : Kind → Type) (ite : InstantiatedTypeEnv TV Γ) => tx TV ite).
-      apply RCanR.
+      apply ACanR.
         apply phoas_extensionality.
         reflexivity.
     apply (Prelude_error "should not be possible").
@@ -145,14 +146,14 @@ Section HaskSkolemizer.
     destruct (eqd_dec ([tx],,take_arg_types_as_tree te) (take_arg_types_as_tree (tx ---> te))).
       rewrite <- e.
       simpl.
-      apply RId.
+      apply AId.
     unfold take_arg_types_as_tree.
       Opaque take_arg_types_as_tree.
       simpl.
       destruct (count_arg_types (te (fun _ : Kind => unit) (ite_unit Γ))).
       simpl.
       replace (tx) with (fun (TV : Kind → Type) (ite : InstantiatedTypeEnv TV Γ) => tx TV ite).
-      apply RuCanR.
+      apply AuCanR.
         apply phoas_extensionality.
         reflexivity.
     apply (Prelude_error "should not be possible").
@@ -231,6 +232,9 @@ Section HaskSkolemizer.
       | RAbsCo   Γ Δ Σ κ σ  σ₁ σ₂  lev => let case_RAbsCo := tt        in _
       | RApp     Γ Δ Σ₁ Σ₂ tx te lev   => let case_RApp := tt          in _
       | RLet     Γ Δ Σ₁ Σ₂ σ₁ σ₂ lev   => let case_RLet := tt          in _
+      | RCut    Γ Δ Σ₁ Σ₁₂ Σ₂ Σ₃ l    => let case_RCut := tt          in _
+      | RLeft   Γ Δ Σ₁ Σ₂  Σ     l    => let case_RLeft := tt in _
+      | RRight  Γ Δ Σ₁ Σ₂  Σ     l    => let case_RRight := tt in _
       | RWhere   Γ Δ Σ₁ Σ₂ Σ₃ σ₁ σ₂ lev => let case_RWhere := tt          in _
       | RJoin    Γ p lri m x q l       => let case_RJoin := tt in _
       | RVoid    _ _           l       => let case_RVoid := tt   in _
@@ -250,7 +254,7 @@ Section HaskSkolemizer.
         apply nd_rule.
         apply SFlat.
         apply RArrange.
-        apply RRight.
+        apply ARight.
         apply d.
 
       destruct case_RBrak.
@@ -287,7 +291,7 @@ Section HaskSkolemizer.
         rewrite H.
         rewrite H0.
         simpl.
-        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; apply RuCanL ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; apply AuCanL ].
         apply nd_rule.
         apply SFlat.
         apply RLit.
@@ -302,7 +306,7 @@ Section HaskSkolemizer.
         rewrite H.
         rewrite H0.
         simpl.
-        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply RuCanR ].
+        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply AuCanR ].
         apply nd_rule.
         apply SFlat.
         apply RVar.
@@ -317,7 +321,7 @@ Section HaskSkolemizer.
         rewrite H.
         rewrite H0.
         simpl.
-        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply RuCanR ].
+        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply AuCanR ].
         apply nd_rule.
         apply SFlat.
         apply RGlobal.
@@ -333,9 +337,9 @@ Section HaskSkolemizer.
         apply nd_rule.
           apply SFlat.
           apply RArrange.
-          eapply RComp.
-          eapply RCossa.
-          eapply RLeft.
+          eapply AComp.
+          eapply AuAssoc.
+          eapply ALeft.
           apply take_arrange.
 
       destruct case_RCast.
@@ -369,14 +373,14 @@ Section HaskSkolemizer.
         rewrite H0.
         simpl.
         eapply nd_comp.
-        eapply nd_prod; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply RCanR ].
+        eapply nd_prod; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ACanR ].
         eapply nd_rule.
         eapply SFlat.
         eapply RArrange.
-        eapply RLeft.
+        eapply ALeft.
         eapply take_unarrange.
 
-        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply RAssoc ].
+        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply AAssoc ].
         eapply nd_rule; eapply SFlat; apply RWhere.
 
       destruct case_RLet.
@@ -392,17 +396,17 @@ Section HaskSkolemizer.
         rewrite H0.
 
         eapply nd_comp.
-        eapply nd_prod; [ eapply nd_rule; eapply SFlat; eapply RArrange; eapply RCanR | eapply nd_id ].
+        eapply nd_prod; [ eapply nd_rule; eapply SFlat; eapply RArrange; eapply ACanR | eapply nd_id ].
 
         set (@RLet Γ Δ Σ₁ (Σ₂,,(take_arg_types_as_tree σ₂ @@@ (h::lev))) σ₁ (drop_arg_types_as_tree σ₂) (h::lev)) as q.
-        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply RAssoc ].
+        eapply nd_comp; [ idtac | eapply nd_rule; apply SFlat; eapply RArrange; apply AAssoc ].
         eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply q ].
         apply nd_prod.
           apply nd_id.
           apply nd_rule.
           eapply SFlat.
           eapply RArrange.
-          eapply RCossa.
+          eapply AuAssoc.
 
       destruct case_RWhere.
         simpl.
@@ -417,16 +421,83 @@ Section HaskSkolemizer.
         rewrite H0.
 
         eapply nd_comp.
-        eapply nd_prod; [ apply nd_id | eapply nd_rule; eapply SFlat; eapply RArrange; eapply RCanR ].
-        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply RAssoc ].
-        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply RLeft; eapply RAssoc ].
+        eapply nd_prod; [ apply nd_id | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ACanR ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AAssoc ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ALeft; eapply AAssoc ].
         eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RWhere ].
         apply nd_prod; [ idtac | eapply nd_id ].
         eapply nd_rule; apply SFlat; eapply RArrange.
-        eapply RComp.
-        eapply RCossa.
-        apply RLeft.
-        eapply RCossa.
+        eapply AComp.
+        eapply AuAssoc.
+        apply ALeft.
+        eapply AuAssoc.
+
+      destruct case_RCut.
+        simpl; destruct l; [ apply nd_rule; apply SFlat; apply RCut | idtac ].
+        set (mapOptionTreeAndFlatten take_arg_types_as_tree Σ₃) as Σ₃''.
+        set (mapOptionTree drop_arg_types_as_tree Σ₃) as Σ₃'''.
+        set (mapOptionTreeAndFlatten take_arg_types_as_tree Σ₁₂) as Σ₁₂''.
+        set (mapOptionTree drop_arg_types_as_tree Σ₁₂) as Σ₁₂'''.
+        destruct (decide_tree_empty Σ₁₂''); [ idtac | apply (Prelude_error "used RCut on a variable with function type") ].
+        destruct (eqd_dec Σ₁₂ Σ₁₂'''); [ idtac | apply (Prelude_error "used RCut on a variable with function type") ].
+        rewrite <- e.
+        eapply nd_comp.
+        eapply nd_prod; [ apply nd_id | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AuAssoc ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AAssoc ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RCut ].
+        apply nd_prod.
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ACanR ].
+        apply nd_rule.
+        apply SFlat.
+        apply RArrange.
+        apply ALeft.
+        destruct s.
+        eapply arrangeCancelEmptyTree with (q:=x).
+        rewrite e0.
+        admit.   (* FIXME, but not serious *)
+        apply nd_id.
+
+      destruct case_RLeft.
+        simpl; destruct l; [ apply nd_rule; apply SFlat; apply RLeft | idtac ].
+        set (mapOptionTreeAndFlatten take_arg_types_as_tree Σ₂) as Σ₂'.
+        set (mapOptionTreeAndFlatten take_arg_types_as_tree Σ) as Σ'.
+        set (mapOptionTree drop_arg_types_as_tree Σ₂) as Σ₂''.
+        set (mapOptionTree drop_arg_types_as_tree Σ) as Σ''.
+        destruct (decide_tree_empty (Σ' @@@ (h::l)));
+          [ idtac | apply (Prelude_error "used RLeft on a variable with function type") ].
+        destruct (eqd_dec Σ Σ''); [ idtac | apply (Prelude_error "used RLeft on a variable with function type") ].
+        rewrite <- e.
+        clear Σ'' e.
+        destruct s.
+        set (arrangeUnCancelEmptyTree _ _ e) as q.
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ALeft; eapply ARight; eapply q ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ALeft; eapply AuCanL; eapply q ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AAssoc ].
+        apply nd_rule.
+        eapply SFlat.
+        eapply RLeft.
+        
+      destruct case_RRight.
+        simpl; destruct l; [ apply nd_rule; apply SFlat; apply RRight | idtac ].
+        set (mapOptionTreeAndFlatten take_arg_types_as_tree Σ₂) as Σ₂'.
+        set (mapOptionTreeAndFlatten take_arg_types_as_tree Σ) as Σ'.
+        set (mapOptionTree drop_arg_types_as_tree Σ₂) as Σ₂''.
+        set (mapOptionTree drop_arg_types_as_tree Σ) as Σ''.
+        destruct (decide_tree_empty (Σ' @@@ (h::l)));
+          [ idtac | apply (Prelude_error "used RRight on a variable with function type") ].
+        destruct (eqd_dec Σ Σ''); [ idtac | apply (Prelude_error "used RRight on a variable with function type") ].
+        rewrite <- e.
+        clear Σ'' e.
+        destruct s.
+        set (arrangeUnCancelEmptyTree _ _ e) as q.
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ALeft; eapply ALeft; eapply q ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ALeft; eapply AuCanR ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AAssoc ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply ALeft; eapply AExch ].  (* yuck *)
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AuAssoc ].
+        eapply nd_rule.
+        eapply SFlat.
+        apply RRight.
 
       destruct case_RVoid.
         simpl.
@@ -434,7 +505,7 @@ Section HaskSkolemizer.
         apply nd_rule.
         apply SFlat.
         apply RVoid.
-        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply RuCanL ].
+        eapply nd_comp; [ idtac | eapply nd_rule; eapply SFlat; eapply RArrange; eapply AuCanL ].
         apply nd_rule.
         apply SFlat.
         apply RVoid.
index ed3cd92..bf51e1a 100644 (file)
@@ -29,55 +29,53 @@ Section HaskStrong.
   }.
   Implicit Arguments StrongCaseBranchWithVVs [[Γ]].
 
-  Inductive Expr : forall Γ (Δ:CoercionEnv Γ), (VV -> LeveledHaskType Γ ★) -> LeveledHaskType Γ ★ -> Type :=
+  Inductive Expr : forall Γ (Δ:CoercionEnv Γ), (VV -> LeveledHaskType Γ ★) -> HaskType Γ ★ -> HaskLevel Γ -> Type :=
 
   (* an "EGlobal" is any variable which is free in the expression which was passed to -fcoqpass (ie bound outside it) *)
-  | EGlobal: forall Γ Δ ξ   (g:Global Γ) v lev,                                                      Expr Γ Δ ξ (g v @@ lev)
+  | EGlobal: forall Γ Δ ξ   (g:Global Γ) v lev,                                                      Expr Γ Δ ξ (g v) lev
 
-  | EVar   : ∀ Γ Δ ξ ev,                                                                             Expr Γ Δ ξ (ξ ev)
-  | ELit   : ∀ Γ Δ ξ lit   l,                                                                        Expr Γ Δ ξ (literalType lit@@l)
-  | EApp   : ∀ Γ Δ ξ t1 t2 l,        Expr Γ Δ ξ (t2--->t1 @@ l)   -> Expr Γ Δ ξ (t2 @@ l)         -> Expr Γ Δ ξ (t1 @@ l)
-  | ELam   : ∀ Γ Δ ξ t1 t2 l ev,              Expr Γ Δ (update_xi ξ l ((ev,t1)::nil)) (t2@@l)      -> Expr Γ Δ ξ (t1--->t2@@l)
-  | ELet   : ∀ Γ Δ ξ tv t  l ev,Expr Γ Δ ξ (tv@@l)->Expr Γ Δ (update_xi ξ l ((ev,tv)::nil))(t@@l)  -> Expr Γ Δ ξ (t@@l)
-  | EEsc   : ∀ Γ Δ ξ ec t  l,     Expr Γ Δ ξ (<[ ec |- t ]> @@ l)                                 -> Expr Γ Δ ξ (t @@ (ec::l))
-  | EBrak  : ∀ Γ Δ ξ ec t  l,     Expr Γ Δ ξ (t @@ (ec::l))                                       -> Expr Γ Δ ξ (<[ ec |- t ]> @@ l)
-  | ECast  : forall Γ Δ ξ t1 t2 (γ:HaskCoercion Γ Δ (t1 ∼∼∼ t2)) l,
-    Expr Γ Δ ξ (t1 @@ l)                        -> Expr Γ Δ ξ (t2 @@ l)
-  | ENote  : ∀ Γ Δ ξ t, Note                      -> Expr Γ Δ ξ t                                 -> Expr Γ Δ ξ t
-  | ETyApp : ∀ Γ Δ κ σ τ ξ l,                    Expr Γ Δ ξ (HaskTAll κ σ @@ l)                   -> Expr Γ Δ ξ (substT σ τ @@ l)
-  | ECoLam : forall Γ Δ κ σ (σ₁ σ₂:HaskType Γ κ) ξ l,
-    Expr Γ (σ₁∼∼∼σ₂::Δ) ξ (σ @@ l)    -> Expr Γ Δ ξ (σ₁∼∼σ₂    ⇒ σ @@ l)
-  | ECoApp : forall Γ Δ κ (σ₁ σ₂:HaskType Γ κ) (γ:HaskCoercion Γ Δ (σ₁∼∼∼σ₂)) σ ξ l,
-    Expr Γ Δ ξ (σ₁ ∼∼ σ₂ ⇒ σ @@ l)            -> Expr Γ Δ ξ (σ        @@l)
+  | EVar   : ∀ Γ Δ ξ ev,                                                                Expr Γ Δ ξ (unlev (ξ ev)) (getlev (ξ ev))
+  | ELit   : ∀ Γ Δ ξ lit   l,                                                                        Expr Γ Δ ξ (literalType lit) l
+  | EApp   : ∀ Γ Δ ξ t1 t2 l,        Expr Γ Δ ξ (t2--->t1) l   -> Expr Γ Δ ξ t2 l         -> Expr Γ Δ ξ (t1) l
+  | ELam   : ∀ Γ Δ ξ t1 t2 l ev,              Expr Γ Δ (update_xi ξ l ((ev,t1)::nil)) t2 l       -> Expr Γ Δ ξ (t1--->t2) l
+  | ELet   : ∀ Γ Δ ξ tv t  l ev,Expr Γ Δ ξ  tv l ->Expr Γ Δ (update_xi ξ l ((ev,tv)::nil)) t l  -> Expr Γ Δ ξ t l
+  | EEsc   : ∀ Γ Δ ξ ec t  l,     Expr Γ Δ ξ  (<[ ec |- t ]>)  l                                  -> Expr Γ Δ ξ t (ec::l)
+  | EBrak  : ∀ Γ Δ ξ ec t  l,     Expr Γ Δ ξ t  (ec::l)                                         -> Expr Γ Δ ξ (<[ ec |- t ]>) l
+  | ECast  : forall Γ Δ ξ t1 t2 (γ:HaskCoercion Γ Δ (t1 ∼∼∼ t2)) l, Expr Γ Δ ξ t1 l             -> Expr Γ Δ ξ t2 l
+  | ENote  : ∀ Γ Δ ξ t l, Note                      -> Expr Γ Δ ξ t l                               -> Expr Γ Δ ξ t l
+  | ETyApp : ∀ Γ Δ κ σ τ ξ l,                    Expr Γ Δ ξ (HaskTAll κ σ) l                   -> Expr Γ Δ ξ (substT σ τ) l
+  | ECoLam : forall Γ Δ κ σ (σ₁ σ₂:HaskType Γ κ) ξ l, Expr Γ (σ₁∼∼∼σ₂::Δ) ξ σ l    -> Expr Γ Δ ξ (σ₁∼∼σ₂    ⇒ σ) l
+  | ECoApp : forall Γ Δ κ (σ₁ σ₂:HaskType Γ κ) (γ:HaskCoercion Γ Δ (σ₁∼∼∼σ₂)) σ ξ l, Expr Γ Δ ξ (σ₁ ∼∼ σ₂ ⇒ σ) l  -> Expr Γ Δ ξ σ l
   | ETyLam : ∀ Γ Δ ξ κ σ l,
-    Expr (κ::Γ) (weakCE Δ) (weakLT○ξ) (HaskTApp (weakF σ) (FreshHaskTyVar _)@@(weakL l))-> Expr Γ Δ ξ (HaskTAll κ σ @@ l)
+    Expr (κ::Γ) (weakCE Δ) (weakLT○ξ) (HaskTApp (weakF σ) (FreshHaskTyVar _)) (weakL l)-> Expr Γ Δ ξ (HaskTAll κ σ) l
   | ECase    : forall Γ Δ ξ l tc tbranches atypes,
-               Expr Γ Δ ξ (caseType tc atypes @@ l) ->
+               Expr Γ Δ ξ (caseType tc atypes) l ->
                Tree ??{ sac : _
                     & { scb : StrongCaseBranchWithVVs tc atypes sac
                     &         Expr (sac_gamma sac Γ)
                                    (sac_delta sac Γ atypes (weakCK'' Δ))
                                    (scbwv_xi scb ξ l)
-                                   (weakLT' (tbranches@@l)) } } ->
-               Expr Γ Δ ξ (tbranches @@ l)
+                                   (weakT' tbranches)
+                                   (weakL' l) } } ->
+               Expr Γ Δ ξ tbranches l
 
   | ELetRec  : ∀ Γ Δ ξ l τ vars,
     distinct (leaves (mapOptionTree (@fst _ _) vars)) ->
     let ξ' := update_xi ξ l (leaves vars) in
-    ELetRecBindings Γ Δ ξ'     l vars ->
-    Expr            Γ Δ ξ' (τ@@l) ->
-    Expr            Γ Δ ξ  (τ@@l)
+    ELetRecBindings Γ Δ ξ'   l vars ->
+    Expr            Γ Δ ξ' τ l ->
+    Expr            Γ Δ ξ  τ l
 
   (* can't avoid having an additional inductive: it is a tree of Expr's, each of whose ξ depends on the type of the entire tree *)
   with ELetRecBindings : ∀ Γ, CoercionEnv Γ -> (VV -> LeveledHaskType Γ ★) -> HaskLevel Γ -> Tree ??(VV*HaskType Γ ★) -> Type :=
   | ELR_nil    : ∀ Γ Δ ξ l  ,                                                                 ELetRecBindings Γ Δ ξ l []
-  | ELR_leaf   : ∀ Γ Δ ξ l v t,                                        Expr Γ Δ ξ (t @@ l) -> ELetRecBindings Γ Δ ξ l [(v,t)]
+  | ELR_leaf   : ∀ Γ Δ ξ l v t,                                        Expr Γ Δ ξ  t  l    -> ELetRecBindings Γ Δ ξ l [(v,t)]
   | ELR_branch : ∀ Γ Δ ξ l t1 t2, ELetRecBindings Γ Δ ξ l t1 -> ELetRecBindings Γ Δ ξ l t2 -> ELetRecBindings Γ Δ ξ l (t1,,t2)
   .
 
   Context {ToStringVV:ToString VV}.
   Context {ToLatexVV:ToLatex VV}.
-  Fixpoint exprToString {Γ}{Δ}{ξ}{τ}(exp:Expr Γ Δ ξ τ) : string :=
+  Fixpoint exprToString {Γ}{Δ}{ξ}{τ}{l}(exp:Expr Γ Δ ξ τ l) : string :=
     match exp with
     | EVar  Γ' _ ξ' ev              => "var."+++ toString ev
     | EGlobal Γ' _ ξ'   g v _       => "global." +++ toString (g:CoreVar)
@@ -87,7 +85,7 @@ Section HaskStrong.
     | EApp  Γ' _ _ _ _ _ e1 e2      => "("+++exprToString e1+++")("+++exprToString e2+++")"
     | EEsc  Γ' _ _ ec t _ e         => "~~("+++exprToString e+++")"
     | EBrak Γ' _ _ ec t _ e         => "<["+++exprToString e+++"]>"
-    | ENote _ _ _ _ n e             => "note."+++exprToString e
+    | ENote _ _ _ _ n _ e           => "note."+++exprToString e
     | ETyApp  Γ Δ κ σ τ ξ l       e => "("+++exprToString e+++")@("+++toString τ+++")"
     | ECoApp Γ Δ κ σ₁ σ₂ γ σ ξ l e  => "("+++exprToString e+++")~(co)"
     | ECast Γ Δ ξ t1 t2 γ l e       => "cast ("+++exprToString e+++"):t2"
@@ -96,7 +94,7 @@ Section HaskStrong.
     | ECase Γ Δ ξ l tc branches atypes escrut alts => "case " +++ exprToString escrut +++ " of FIXME"
     | ELetRec _ _ _ _ _ vars vu elrb e => "letrec FIXME in " +++ exprToString e
     end.
-  Instance ExprToString Γ Δ ξ τ : ToString (Expr Γ Δ ξ τ) := { toString := exprToString }.
+  Instance ExprToString Γ Δ ξ τ l : ToString (Expr Γ Δ ξ τ l) := { toString := exprToString }.
 
 End HaskStrong.
 Implicit Arguments StrongCaseBranchWithVVs [[Γ]].
index 791fdf5..1b9f6af 100644 (file)
@@ -6,6 +6,7 @@ Generalizable All Variables.
 Require Import Preamble.
 Require Import General.
 Require Import NaturalDeduction.
+Require Import NaturalDeductionContext.
 Require Import Coq.Strings.String.
 Require Import Coq.Lists.List.
 Require Import Coq.Init.Specif.
@@ -15,20 +16,7 @@ Require Import HaskStrong.
 Require Import HaskProof.
 
 Section HaskStrongToProof.
-
-Definition pivotContext {T} a b c : @Arrange T ((a,,b),,c) ((a,,c),,b) :=
-  RComp (RComp (RCossa _ _ _) (RLeft a (RExch c b))) (RAssoc _ _ _).
-
-Definition pivotContext' {T} a b c : @Arrange T (a,,(b,,c)) (b,,(a,,c)) :=
-  RComp (RComp (RAssoc _ _ _) (RRight c (RExch b a))) (RCossa _ _ _).
-
-Definition copyAndPivotContext {T} a b c : @Arrange T ((a,,b),,(c,,b)) ((a,,c),,b).
-  eapply RComp; [ idtac | apply (RLeft (a,,c) (RCont b)) ].
-  eapply RComp; [ idtac | apply RCossa ]. 
-  eapply RComp; [ idtac | apply (RRight b (pivotContext a b c)) ].
-  apply RAssoc.
-  Defined.
-  
 Context {VV:Type}{eqd_vv:EqDecidable VV}.
 
 (* maintenance of Xi *)
@@ -520,8 +508,8 @@ Lemma update_xiv_lemma `{EQD_VV:EqDecidable VV} Γ ξ (lev:HaskLevel Γ)(varstyp
   apply q.
   Qed.
 
-Fixpoint expr2antecedent {Γ'}{Δ'}{ξ'}{τ'}(exp:Expr Γ' Δ' ξ' τ') : Tree ??VV :=
-  match exp as E in Expr Γ Δ ξ τ with
+Fixpoint expr2antecedent {Γ'}{Δ'}{ξ'}{τ'}{l'}(exp:Expr Γ' Δ' ξ' τ' l') : Tree ??VV :=
+  match exp as E in Expr Γ Δ ξ τ l with
   | EGlobal  Γ Δ ξ _ _ _                          => []
   | EVar     Γ Δ ξ ev                             => [ev]
   | ELit     Γ Δ ξ lit lev                        => []
@@ -531,7 +519,7 @@ Fixpoint expr2antecedent {Γ'}{Δ'}{ξ'}{τ'}(exp:Expr Γ' Δ' ξ' τ') : Tree ?
   | EEsc     Γ Δ ξ ec t lev e                     => expr2antecedent e
   | EBrak    Γ Δ ξ ec t lev e                     => expr2antecedent e
   | ECast    Γ Δ ξ γ t1 t2 lev      e             => expr2antecedent e
-  | ENote    Γ Δ ξ t n e                          => expr2antecedent e
+  | ENote    Γ Δ ξ t l n e                        => expr2antecedent e
   | ETyLam   Γ Δ ξ κ σ l e                        => expr2antecedent e
   | ECoLam   Γ Δ κ σ σ₁ σ₂ ξ l             e      => expr2antecedent e
   | ECoApp   Γ Δ κ γ σ₁ σ₂ σ ξ l      e           => expr2antecedent e
@@ -546,7 +534,7 @@ Fixpoint expr2antecedent {Γ'}{Δ'}{ξ'}{τ'}(exp:Expr Γ' Δ' ξ' τ') : Tree ?
                             & Expr (sac_gamma sac Γ)
                                    (sac_delta sac Γ atypes (weakCK'' Δ))
                                    (scbwv_xi scb ξ l)
-                                   (weakLT' (tbranches@@l)) } }
+                                   (weakT' tbranches) (weakL' l)} }
       ): Tree ??VV :=
       match alts with
         | T_Leaf None     => []
@@ -566,7 +554,7 @@ Definition mkProofCaseBranch {Γ}{Δ}{ξ}{l}{tc}{tbranches}{atypes}
                             & Expr (sac_gamma sac Γ)
                                    (sac_delta sac Γ atypes (weakCK'' Δ))
                                    (scbwv_xi scb ξ l)
-                                   (weakLT' (tbranches@@l)) } })
+                                   (weakT' tbranches) (weakL' l) } })
   : { sac : _ & ProofCaseBranch tc Γ Δ l tbranches atypes sac }.
   destruct alt.
   exists x.
@@ -633,16 +621,16 @@ Definition factorContextLeft
             (* where the leaf is v *)
             apply inr.
               subst.
-              apply RuCanR.
+              apply AuCanR.
 
             (* where the leaf is NOT v *)
             apply inl.
-              apply RuCanL.
+              apply AuCanL.
   
         (* empty leaf *)
         destruct case_None.
           apply inl; simpl in *.
-          apply RuCanR.
+          apply AuCanR.
   
       (* branch *)
       refine (
@@ -662,45 +650,45 @@ Definition factorContextLeft
     destruct case_Neither.
       apply inl.
       simpl.
-      eapply RComp; [idtac | apply RuCanL ].
-        exact (RComp
+      eapply AComp; [idtac | apply AuCanL ].
+        exact (AComp
           (* order will not matter because these are central as morphisms *)
-          (RRight _ (RComp lpf (RCanL _)))
-          (RLeft  _ (RComp rpf (RCanL _)))).
+          (ARight _ (AComp lpf (ACanL _)))
+          (ALeft  _ (AComp rpf (ACanL _)))).
 
     destruct case_Right.
       apply inr.
       unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *.
       fold (stripOutVars (v::nil)).
-      eapply RComp; [ idtac | eapply pivotContext' ].
-      eapply RComp.
-      eapply RRight.
-      eapply RComp.
+      eapply AComp; [ idtac | eapply pivotContext' ].
+      eapply AComp.
+      eapply ARight.
+      eapply AComp.
       apply lpf.
-      apply RCanL.
-      eapply RLeft.
+      apply ACanL.
+      eapply ALeft.
       apply rpf.
 
     destruct case_Left.
       apply inr.
       fold (stripOutVars (v::nil)).
       simpl.
-      eapply RComp.
-      eapply RLeft.
-      eapply RComp.
+      eapply AComp.
+      eapply ALeft.
+      eapply AComp.
       apply rpf.
       simpl.
-      eapply RCanL.
-      eapply RComp; [ idtac | eapply RCossa ].
-      eapply RRight.
+      eapply ACanL.
+      eapply AComp; [ idtac | eapply AuAssoc ].
+      eapply ARight.
       apply lpf.
 
     destruct case_Both.
       apply inr.
       simpl.
-      eapply RComp; [ idtac | eapply RRight; eapply RCont ].
-      eapply RComp; [ eapply RRight; eapply lpf | idtac ].
-      eapply RComp; [ eapply RLeft; eapply rpf | idtac ].
+      eapply AComp; [ idtac | eapply ARight; eapply ACont ].
+      eapply AComp; [ eapply ARight; eapply lpf | idtac ].
+      eapply AComp; [ eapply ALeft; eapply rpf | idtac ].
       clear lpf rpf.
       simpl.
       apply arrangeSwapMiddle.
@@ -738,16 +726,16 @@ Definition factorContextRight
             (* where the leaf is v *)
             apply inr.
               subst.
-              apply RuCanL.
+              apply AuCanL.
 
             (* where the leaf is NOT v *)
             apply inl.
-              apply RuCanR.
+              apply AuCanR.
   
         (* empty leaf *)
         destruct case_None.
           apply inl; simpl in *.
-          apply RuCanR.
+          apply AuCanR.
   
       (* branch *)
       refine (
@@ -766,51 +754,51 @@ Definition factorContextRight
 
     destruct case_Neither.
       apply inl.
-      eapply RComp; [idtac | apply RuCanR ].
-        exact (RComp
+      eapply AComp; [idtac | apply AuCanR ].
+        exact (AComp
           (* order will not matter because these are central as morphisms *)
-          (RRight _ (RComp lpf (RCanR _)))
-          (RLeft  _ (RComp rpf (RCanR _)))).
+          (ARight _ (AComp lpf (ACanR _)))
+          (ALeft  _ (AComp rpf (ACanR _)))).
 
 
     destruct case_Right.
       apply inr.
       fold (stripOutVars (v::nil)).
-      set (RRight (mapOptionTree ξ ctx2)  (RComp lpf ((RCanR _)))) as q.
+      set (ARight (mapOptionTree ξ ctx2)  (AComp lpf ((ACanR _)))) as q.
       simpl in *.
-      eapply RComp.
+      eapply AComp.
       apply q.
       clear q.
       clear lpf.
       unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *.
-      eapply RComp; [ idtac | apply RAssoc ].
-      apply RLeft.
+      eapply AComp; [ idtac | apply AAssoc ].
+      apply ALeft.
       apply rpf.
 
     destruct case_Left.
       apply inr.
       unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *.
       fold (stripOutVars (v::nil)).
-      eapply RComp; [ idtac | eapply pivotContext ].
-      set (RComp rpf (RCanR _ )) as rpf'.
-      set (RLeft ((mapOptionTree ξ (stripOutVars (v :: nil) ctx1),, [ξ v])) rpf') as qq.
+      eapply AComp; [ idtac | eapply pivotContext ].
+      set (AComp rpf (ACanR _ )) as rpf'.
+      set (ALeft ((mapOptionTree ξ (stripOutVars (v :: nil) ctx1),, [ξ v])) rpf') as qq.
       simpl in *.
-      eapply RComp; [ idtac | apply qq ].
+      eapply AComp; [ idtac | apply qq ].
       clear qq rpf' rpf.
-      apply (RRight (mapOptionTree ξ ctx2)).
+      apply (ARight (mapOptionTree ξ ctx2)).
       apply lpf.
 
     destruct case_Both.
       apply inr.
       unfold mapOptionTree in *; simpl; fold (mapOptionTree ξ) in *.
       fold (stripOutVars (v::nil)).
-      eapply RComp; [ idtac | eapply copyAndPivotContext ].
+      eapply AComp; [ idtac | eapply copyAndPivotContext ].
         (* order will not matter because these are central as morphisms *)
-        exact (RComp (RRight _ lpf) (RLeft _ rpf)).
+        exact (AComp (ARight _ lpf) (ALeft _ rpf)).
 
     Defined.
 
-(* same as before, but use RWeak if necessary *)
+(* same as before, but use AWeak if necessary *)
 Definition factorContextLeftAndWeaken  
      (Γ:TypeEnv)(Δ:CoercionEnv Γ)
      v      (* variable to be pivoted, if found *)
@@ -821,8 +809,8 @@ Definition factorContextLeftAndWeaken
           (mapOptionTree ξ ([v],,(stripOutVars (v::nil) ctx))        ).
   set (factorContextLeft Γ Δ v ctx ξ) as q.
   destruct q; auto.
-  eapply RComp; [ apply a | idtac ].
-  refine (RRight _ (RWeak _)).
+  eapply AComp; [ apply a | idtac ].
+  refine (ARight _ (AWeak _)).
   Defined.
 
 Definition factorContextLeftAndWeaken''
@@ -848,7 +836,7 @@ Definition factorContextLeftAndWeaken''
   unfold mapOptionTree; simpl in *.
     intros.
     rewrite (@stripping_nothing_is_inert Γ); auto.
-    apply RuCanL.
+    apply AuCanL.
     intros.
     unfold mapOptionTree in *.
     simpl in *.
@@ -867,15 +855,15 @@ Definition factorContextLeftAndWeaken''
       unfold stripOutVars in q.
       rewrite q in IHv1'.
       clear q.
-    eapply RComp; [ idtac | eapply RAssoc ].
-    eapply RComp; [ idtac | eapply IHv1' ].
+    eapply AComp; [ idtac | eapply AAssoc ].
+    eapply AComp; [ idtac | eapply IHv1' ].
     clear IHv1'.
     apply IHv2; auto.
     auto.
     auto.
     Defined.
 
-(* same as before, but use RWeak if necessary *)
+(* same as before, but use AWeak if necessary *)
 Definition factorContextRightAndWeaken  
      (Γ:TypeEnv)(Δ:CoercionEnv Γ)
      v      (* variable to be pivoted, if found *)
@@ -886,8 +874,8 @@ Definition factorContextRightAndWeaken
           (mapOptionTree ξ ((stripOutVars (v::nil) ctx),,[v])        ).
   set (factorContextRight Γ Δ v ctx ξ) as q.
   destruct q; auto.
-  eapply RComp; [ apply a | idtac ].
-  refine (RLeft _ (RWeak _)).
+  eapply AComp; [ apply a | idtac ].
+  refine (ALeft _ (AWeak _)).
   Defined.
 
 Definition factorContextRightAndWeaken''
@@ -911,7 +899,7 @@ Definition factorContextRightAndWeaken''
   unfold mapOptionTree; simpl in *.
     intros.
     rewrite (@stripping_nothing_is_inert Γ); auto.
-    apply RuCanR.
+    apply AuCanR.
     intros.
     unfold mapOptionTree in *.
     simpl in *.
@@ -930,14 +918,14 @@ Definition factorContextRightAndWeaken''
     fold X in IHv2'.
     set (distinct_app _ _ _ H) as H'.
     destruct H' as [H1 H2].
-    set (RComp (IHv1 _ H1) (IHv2' H2)) as qq.
-    eapply RComp.
+    set (AComp (IHv1 _ H1) (IHv2' H2)) as qq.
+    eapply AComp.
       apply qq.
       clear qq IHv2' IHv2 IHv1.
       rewrite strip_swap_lemma.
       rewrite strip_twice_lemma.
       rewrite (notin_strip_inert' v1 (leaves v2)).
-        apply RCossa.
+        apply AuAssoc.
         apply distinct_swap.
         auto.
     Defined.
@@ -1067,7 +1055,7 @@ Lemma case_lemma : forall Γ Δ ξ l tc tbranches atypes e
             ??{sac : StrongAltCon &
               {scb : StrongCaseBranchWithVVs VV eqd_vv tc atypes sac &
               Expr (sac_gamma sac Γ) (sac_delta sac Γ atypes (weakCK'' Δ))
-                (scbwv_xi scb ξ l) (weakLT' (tbranches @@  l))}}),
+                (scbwv_xi scb ξ l) (weakT' tbranches) (weakL' l)}}),
 
       (mapOptionTreeAndFlatten (fun x => pcb_freevars (projT2 x))
         (mapOptionTree mkProofCaseBranch alts'))
@@ -1094,40 +1082,40 @@ Lemma case_lemma : forall Γ Δ ξ l tc tbranches atypes e
   Qed.
 
 Definition expr2proof  :
-  forall Γ Δ ξ τ (e:Expr Γ Δ ξ τ),
-    ND Rule [] [Γ > Δ > mapOptionTree ξ (expr2antecedent e) |- [unlev τ] @ getlev τ].
+  forall Γ Δ ξ τ l (e:Expr Γ Δ ξ τ l),
+    ND Rule [] [Γ > Δ > mapOptionTree ξ (expr2antecedent e) |- [τ] @ l].
 
-  refine (fix expr2proof Γ' Δ' ξ' τ' (exp:Expr Γ' Δ' ξ' τ') {struct exp}
-    : ND Rule [] [Γ' > Δ' > mapOptionTree ξ' (expr2antecedent exp) |- [unlev τ'] @ getlev τ'] :=
-    match exp as E in Expr Γ Δ ξ τ with
+  refine (fix expr2proof Γ' Δ' ξ' τ' l' (exp:Expr Γ' Δ' ξ' τ' l') {struct exp}
+    : ND Rule [] [Γ' > Δ' > mapOptionTree ξ' (expr2antecedent exp) |- [τ'] @ l'] :=
+    match exp as E in Expr Γ Δ ξ τ l with
     | EGlobal  Γ Δ ξ     g v lev                    => let case_EGlobal := tt in _
     | EVar     Γ Δ ξ ev                             => let case_EVar := tt in _
     | ELit     Γ Δ ξ lit lev                        => let case_ELit := tt in _
     | EApp     Γ Δ ξ t1 t2 lev e1 e2                => let case_EApp := tt in 
-                                                        (fun e1' e2' => _) (expr2proof _ _ _ _ e1) (expr2proof _ _ _ _ e2)
-    | ELam     Γ Δ ξ t1 t2 lev v    e               => let case_ELam := tt in (fun e' => _) (expr2proof _ _ _ _ e)
+                                                        (fun e1' e2' => _) (expr2proof _ _ _ _ _ e1) (expr2proof _ _ _ _ _ e2)
+    | ELam     Γ Δ ξ t1 t2 lev v    e               => let case_ELam := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
     | ELet     Γ Δ ξ tv t      v lev ev ebody       => let case_ELet := tt in 
-                                                       (fun pf_let pf_body => _) (expr2proof _ _ _ _ ev) (expr2proof _ _ _ _ ebody) 
+                                                       (fun pf_let pf_body => _) (expr2proof _ _ _ _ _ ev) (expr2proof _ _ _ _ _ ebody) 
     | ELetRec  Γ Δ ξ lev t tree disti branches ebody =>
       let ξ' := update_xi ξ lev (leaves tree) in
-      let case_ELetRec := tt in  (fun e' subproofs => _) (expr2proof _ _ _ _ ebody) 
+      let case_ELetRec := tt in  (fun e' subproofs => _) (expr2proof _ _ _ _ _ ebody) 
         ((fix subproofs Γ'' Δ'' ξ'' lev'' (tree':Tree ??(VV * HaskType Γ'' ★))
         (branches':ELetRecBindings Γ'' Δ'' ξ'' lev'' tree')
         : LetRecSubproofs Γ'' Δ'' ξ'' lev'' tree' branches' :=
         match branches' as B in ELetRecBindings G D X L T return LetRecSubproofs G D X L T B with
           | ELR_nil    Γ Δ ξ lev  => lrsp_nil _ _ _ _
-          | ELR_leaf   Γ Δ ξ l v t e => lrsp_leaf Γ Δ ξ l v t e (expr2proof _ _ _ _ e)
+          | ELR_leaf   Γ Δ ξ l v t e => lrsp_leaf Γ Δ ξ l v t e (expr2proof _ _ _ _ _ e)
           | ELR_branch Γ Δ ξ lev t1 t2 b1 b2 => lrsp_cons _ _ _ _ _ _ _ _ (subproofs _ _ _ _ _ b1) (subproofs _ _ _ _ _ b2)
         end
         ) _ _ _ _ tree branches)
-    | EEsc     Γ Δ ξ ec t lev e                     => let case_EEsc    := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | EBrak    Γ Δ ξ ec t lev e                     => let case_EBrak   := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | ECast    Γ Δ ξ γ t1 t2 lev      e             => let case_ECast   := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | ENote    Γ Δ ξ t n e                          => let case_ENote   := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | ETyLam   Γ Δ ξ κ σ l e                        => let case_ETyLam  := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | ECoLam   Γ Δ κ σ σ₁ σ₂ ξ l             e      => let case_ECoLam  := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | ECoApp   Γ Δ κ σ₁ σ₂ σ γ ξ l      e           => let case_ECoApp  := tt in (fun e' => _) (expr2proof _ _ _ _ e)
-    | ETyApp   Γ Δ κ σ τ ξ  l    e                  => let case_ETyApp  := tt in (fun e' => _) (expr2proof _ _ _ _ e)
+    | EEsc     Γ Δ ξ ec t lev e                     => let case_EEsc    := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | EBrak    Γ Δ ξ ec t lev e                     => let case_EBrak   := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | ECast    Γ Δ ξ γ t1 t2 lev      e             => let case_ECast   := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | ENote    Γ Δ ξ t _ n e                        => let case_ENote   := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | ETyLam   Γ Δ ξ κ σ l e                        => let case_ETyLam  := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | ECoLam   Γ Δ κ σ σ₁ σ₂ ξ l             e      => let case_ECoLam  := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | ECoApp   Γ Δ κ σ₁ σ₂ σ γ ξ l      e           => let case_ECoApp  := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
+    | ETyApp   Γ Δ κ σ τ ξ  l    e                  => let case_ETyApp  := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
     | ECase    Γ Δ ξ l tc tbranches atypes e alts' => 
       let dcsp :=
         ((fix mkdcsp (alts:
@@ -1135,7 +1123,7 @@ Definition expr2proof  :
                             & Expr (sac_gamma sac Γ)
                                    (sac_delta sac Γ atypes (weakCK'' Δ))
                                    (scbwv_xi scb ξ l)
-                                   (weakLT' (tbranches@@l)) } })
+                                   (weakT' tbranches) (weakL' l) } })
           : ND Rule [] (mapOptionTree (fun x => pcb_judg (projT2 (mkProofCaseBranch x))) alts) :=
         match alts as ALTS return ND Rule [] 
           (mapOptionTree (fun x => pcb_judg (projT2 (mkProofCaseBranch x))) ALTS) with
@@ -1144,12 +1132,12 @@ Definition expr2proof  :
           | T_Leaf (Some x)  =>
             match x as X return ND Rule [] [pcb_judg (projT2 (mkProofCaseBranch X))] with
             existT sac (existT scbx ex) =>
-            (fun e' => let case_leaf := tt in _) (expr2proof _ _ _ _ ex)
+            (fun e' => let case_leaf := tt in _) (expr2proof _ _ _ _ _ ex)
         end
         end) alts')
-        in let case_ECase := tt in (fun e' => _) (expr2proof _ _ _ _ e)
+        in let case_ECase := tt in (fun e' => _) (expr2proof _ _ _ _ _ e)
     end
-  ); clear exp ξ' τ' Γ' Δ' expr2proof; try clear mkdcsp.
+  ); clear exp ξ' τ' Γ' Δ' l' expr2proof; try clear mkdcsp.
 
     destruct case_EGlobal.
       apply nd_rule.
@@ -1228,7 +1216,6 @@ Definition expr2proof  :
       auto.
 
     destruct case_ENote.
-      destruct t.
       eapply nd_comp; [ idtac | eapply nd_rule; apply RNote ].
       apply e'.
       auto.
index d86d05b..deb7adc 100644 (file)
@@ -107,10 +107,10 @@ Section HaskStrongToWeak.
       | (vv,wev)::rest => update_chi (update_chi' χ rest) vv wev
     end.
 
-  Fixpoint exprToWeakExpr {Γ}{Δ}{ξ}{τ}(χ:VV->???WeakExprVar)(exp:@Expr _ eqVV Γ Δ ξ τ)
+  Fixpoint exprToWeakExpr {Γ}{Δ}{ξ}{τ}{l}(χ:VV->???WeakExprVar)(exp:@Expr _ eqVV Γ Δ ξ τ l)
     : InstantiatedTypeEnv (fun _ => WeakTypeVar) Γ
     -> UniqM WeakExpr :=
-    match exp as E in @Expr _ _ G D X L return InstantiatedTypeEnv (fun _ => WeakTypeVar) G -> UniqM WeakExpr with
+    match exp as E in @Expr _ _ G D X T L return InstantiatedTypeEnv (fun _ => WeakTypeVar) G -> UniqM WeakExpr with
     | EVar  Γ' _ ξ' ev              => fun ite => match χ ev with OK v => return WEVar v | Error s => failM s end
     | EGlobal Γ' _ ξ'   g v lev     => fun ite => bind tv' = mapM (ilist_to_list (ilmap (fun κ x => typeToWeakType x ite) v))
                                                 ; return (fold_left (fun x y => WETyApp x y) tv' (WEVar g))
@@ -133,7 +133,7 @@ Section HaskStrongToWeak.
     | EBrak Γ' _ _ ec t _ e         => fun ite => bind t' = typeToWeakType t ite
                                                 ; bind e' = exprToWeakExpr χ e ite
                                                 ; return WEBrak hetmet_brak (ec _ ite) e' t'
-    | ENote _ _ _ _ n e             => fun ite => bind e' = exprToWeakExpr χ e ite
+    | ENote _ _ _ _ _ n e           => fun ite => bind e' = exprToWeakExpr χ e ite
                                                 ; return WENote n e'
     | ETyApp  Γ Δ κ σ τ ξ l       e => fun ite => bind t' = typeToWeakType τ ite
                                                 ; bind e' = exprToWeakExpr χ e ite
@@ -158,7 +158,7 @@ Section HaskStrongToWeak.
                   ; bind tbranches' = @typeToWeakType Γ _ tbranches ite
                   ; bind escrut'    = exprToWeakExpr χ escrut ite
                   ; bind branches'  =
-                      ((fix caseBranches (tree:Tree ??{sac : _ & { scb : StrongCaseBranchWithVVs VV _ _ _ sac & Expr _ _ _ _ } })
+                      ((fix caseBranches (tree:Tree ??{sac : _ & { scb : StrongCaseBranchWithVVs VV _ _ _ sac & Expr _ _ _ _ _ } })
                             : UniqM (Tree ??(WeakAltCon*list WeakTypeVar*list WeakCoerVar*list WeakExprVar*WeakExpr)) :=
                             match tree with
                               | T_Leaf None           => return []
@@ -208,7 +208,7 @@ Section HaskStrongToWeak.
   end.
 
 
-  Fixpoint strongExprToWeakExpr (us:UniqSupply){Γ}{Δ}{ξ}{τ}(exp:@Expr _ eqVV Γ Δ ξ τ)
+  Fixpoint strongExprToWeakExpr (us:UniqSupply){Γ}{Δ}{ξ}{τ}{l}(exp:@Expr _ eqVV Γ Δ ξ τ l)
     (ite:InstantiatedTypeEnv (fun _ => WeakTypeVar) Γ)
     : ???WeakExpr :=
     match exprToWeakExpr (fun v => Error ("unbound variable " +++ toString v)) exp ite with
index 670d8bd..077f7b5 100644 (file)
@@ -108,4 +108,4 @@ Fixpoint simplifyWeakExpr (me:WeakExpr) : WeakExpr :=
   (* un-letrec-ify multi branch letrecs *)
   | WELetRec mlr e                       => WELetRec mlr (simplifyWeakExpr e )
   end.
-*)
\ No newline at end of file
+*)
index 1565637..156c2ce 100644 (file)
@@ -399,12 +399,10 @@ Definition weakPsi {Γ}{Δ:CoercionEnv Γ} {κ}(ψ:WeakCoerVar -> ???(HaskCoVar
   Defined.
 
 (* attempt to "cast" an expression by simply checking if it already had the desired type, and failing otherwise *)
-Definition castExpr (we:WeakExpr)(err_msg:string) {Γ} {Δ} {ξ} {τ} τ' (e:@Expr _ CoreVarEqDecidable Γ Δ ξ τ)
-  : ???(@Expr _ CoreVarEqDecidable Γ Δ ξ τ').
+Definition castExpr (we:WeakExpr)(err_msg:string) {Γ} {Δ} {ξ} {τ} {l} τ' l' (e:@Expr _ CoreVarEqDecidable Γ Δ ξ τ l)
+  : ???(@Expr _ CoreVarEqDecidable Γ Δ ξ τ' l').
   apply (addErrorMessage ("castExpr " +++ err_msg)).
   intros.
-  destruct τ  as [τ  l].
-  destruct τ' as [τ' l'].
   destruct (eqd_dec l l'); [ idtac
     | apply (Error ("level mismatch in castExpr, invoked by "+++err_msg+++eol+++
                     "  got: " +++(fold_left (fun x y => y+++","+++y) (map (toString ○ haskTyVarToType) l) "")+++eol+++
@@ -543,7 +541,7 @@ Definition weakExprToStrongExpr : forall
     (ig:CoreVar -> bool)
     (τ:HaskType Γ ★)
     (lev:HaskLevel Γ),
-    WeakExpr -> ???(@Expr _ CoreVarEqDecidable Γ Δ ξ (τ @@ lev) ).
+    WeakExpr -> ???(@Expr _ CoreVarEqDecidable Γ Δ ξ τ lev ).
   refine ((
     fix weakExprToStrongExpr 
     (Γ:TypeEnv)
@@ -554,15 +552,15 @@ Definition weakExprToStrongExpr : forall
     (ig:CoreVar -> bool)
     (τ:HaskType Γ ★)
     (lev:HaskLevel Γ)
-    (we:WeakExpr) : ???(@Expr _ CoreVarEqDecidable Γ Δ ξ (τ @@ lev) )  :=
+    (we:WeakExpr) : ???(@Expr _ CoreVarEqDecidable Γ Δ ξ τ lev )  :=
     addErrorMessage ("in weakExprToStrongExpr " +++ toString we)
     match we with
 
     | WEVar   v                         => if ig v
-                                              then OK ((EGlobal Γ Δ ξ (mkGlobal Γ τ v) INil lev) : Expr Γ Δ ξ (τ@@lev))
-                                              else castExpr we ("WEVar "+++toString (v:CoreVar)) (τ @@ lev) (EVar Γ Δ ξ v)
+                                              then OK ((EGlobal Γ Δ ξ (mkGlobal Γ τ v) INil lev) : Expr Γ Δ ξ τ lev)
+                                              else castExpr we ("WEVar "+++toString (v:CoreVar)) τ lev (EVar Γ Δ ξ v)
 
-    | WELit   lit                       => castExpr we ("WELit "+++toString lit) (τ @@ lev) (ELit Γ Δ ξ lit lev)
+    | WELit   lit                       => castExpr we ("WELit "+++toString lit) τ lev (ELit Γ Δ ξ lit lev)
 
     | WELam   ev ebody                  => weakTypeToTypeOfKind φ ev ★ >>= fun tv =>
                                              weakTypeOfWeakExpr ebody >>= fun tbody =>
@@ -570,24 +568,24 @@ Definition weakExprToStrongExpr : forall
                                                  let ξ' := update_xi ξ lev (((ev:CoreVar),tv)::nil) in
                                                  let ig' := update_ig ig ((ev:CoreVar)::nil) in
                                                    weakExprToStrongExpr Γ Δ φ ψ ξ' ig' tbody' lev ebody >>= fun ebody' =>
-                                                     castExpr we "WELam" (τ@@lev) (ELam Γ Δ ξ tv tbody' lev ev ebody')
+                                                     castExpr we "WELam" τ lev (ELam Γ Δ ξ tv tbody' lev ev ebody')
 
     | WEBrak  _ ec e tbody              => φ (`ec) >>= fun ec' =>
                                              weakTypeToTypeOfKind φ tbody ★ >>= fun tbody' =>
                                                weakExprToStrongExpr Γ Δ φ ψ ξ ig tbody' ((ec')::lev) e >>= fun e' =>
-                                                 castExpr we "WEBrak" (τ@@lev) (EBrak Γ Δ ξ ec' tbody' lev e')
+                                                 castExpr we "WEBrak" τ lev (EBrak Γ Δ ξ ec' tbody' lev e')
 
     | WEEsc   _ ec e tbody              => φ ec >>= fun ec'' =>
                                            weakTypeToTypeOfKind φ tbody ★ >>= fun tbody' =>
                                            match lev with
                                              | nil       => Error "ill-leveled escapification"
                                              | ec'::lev' => weakExprToStrongExpr Γ Δ φ ψ ξ ig (<[ ec' |- tbody' ]>) lev' e
-                                               >>= fun e' => castExpr we "WEEsc" (τ@@lev) (EEsc Γ Δ ξ ec' tbody' lev' e')
+                                               >>= fun e' => castExpr we "WEEsc" τ lev (EEsc Γ Δ ξ ec' tbody' lev' e')
                                            end
 
     | WECSP   _ ec e tbody              => Error "FIXME: CSP not supported beyond HaskWeak stage"
 
-    | WENote  n e                       => weakExprToStrongExpr Γ Δ φ ψ ξ ig τ lev e >>= fun e' => OK (ENote _ _ _ _ n e')
+    | WENote  n e                       => weakExprToStrongExpr Γ Δ φ ψ ξ ig τ lev e >>= fun e' => OK (ENote _ _ _ _ _ n e')
 
     | WELet   v ve  ebody               => weakTypeToTypeOfKind φ v ★  >>= fun tv =>
                                              weakExprToStrongExpr Γ Δ φ ψ ξ ig tv lev ve >>= fun ve' =>
@@ -607,7 +605,7 @@ Definition weakExprToStrongExpr : forall
                                                weakTypeToTypeOfKind φ2 te ★ >>= fun τ' =>
                                                  weakExprToStrongExpr _ (weakCE Δ) φ2
                                                    (fun x => (ψ x) >>= fun y => OK (weakCV y)) (weakLT○ξ) ig _ (weakL lev) e
-                                                     >>= fun e' => castExpr we "WETyLam2" _ (ETyLam Γ Δ ξ tv (mkTAll' τ') lev e')
+                                                     >>= fun e' => castExpr we "WETyLam2" _ _ (ETyLam Γ Δ ξ tv (mkTAll' τ') lev e')
 
     | WETyApp e t                       => weakTypeOfWeakExpr e >>= fun te =>
                                            match te with
@@ -616,7 +614,7 @@ Definition weakExprToStrongExpr : forall
                                                  weakTypeToTypeOfKind φ2 te' ★ >>= fun te'' =>
                                                    weakExprToStrongExpr Γ Δ φ ψ ξ ig (mkTAll te'') lev e >>= fun e' =>
                                                      weakTypeToTypeOfKind φ t (wtv:Kind) >>= fun t' =>
-                                                       castExpr we "WETyApp" _ (ETyApp Γ Δ wtv (mkTAll' te'') t' ξ lev e')
+                                                       castExpr we "WETyApp" _ _ (ETyApp Γ Δ wtv (mkTAll' te'') t' ξ lev e')
                                              | _                 => Error ("weakTypeToType: WETyApp body with type "+++toString te)
                                            end
 
@@ -629,7 +627,7 @@ Definition weakExprToStrongExpr : forall
                                                    weakTypeToTypeOfKind φ t2 κ >>= fun t2'' =>
                                                      weakTypeToTypeOfKind φ t3 ★ >>= fun t3'' =>
                                                        weakExprToStrongExpr Γ Δ φ ψ ξ ig (t1'' ∼∼ t2'' ⇒ τ) lev e >>= fun e' =>
-                                                         castExpr we "WECoApp" _ e' >>= fun e'' =>
+                                                         castExpr we "WECoApp" _ _ e' >>= fun e'' =>
                                                            OK (ECoApp Γ Δ κ t1'' t2''
                                                              (weakCoercionToHaskCoercion _ _ _ co) τ ξ lev e'')
                                                  end
@@ -642,13 +640,13 @@ Definition weakExprToStrongExpr : forall
                                                weakTypeToTypeOfKind φ t1 cv >>= fun t1' =>
                                                  weakTypeToTypeOfKind φ t2 cv >>= fun t2' =>
                                                    weakExprToStrongExpr Γ (_ :: Δ) φ (weakPsi ψ) ξ ig te' lev e >>= fun e' =>
-                                                     castExpr we "WECoLam" _ (ECoLam Γ Δ cv te' t1' t2' ξ lev e')
+                                                     castExpr we "WECoLam" _ _ (ECoLam Γ Δ cv te' t1' t2' ξ lev e')
 
     | WECast  e co                      => let (t1,t2) := weakCoercionTypes co in
                                              weakTypeToTypeOfKind φ t1 ★ >>= fun t1' =>
                                                weakTypeToTypeOfKind φ t2 ★ >>= fun t2' =>
                                                    weakExprToStrongExpr Γ Δ φ ψ ξ ig t1' lev e >>= fun e' =>
-                                                     castExpr we "WECast" _ 
+                                                     castExpr we "WECast" _ _
                                                        (ECast Γ Δ ξ t1' t2' (weakCoercionToHaskCoercion _ _ _ co) lev e')
 
     | WELetRec rb   e                   =>
@@ -679,7 +677,7 @@ Definition weakExprToStrongExpr : forall
                 weakTypeToTypeOfKind φ tbranches ★  >>= fun tbranches' =>
                   (fix mkTree (t:Tree ??(WeakAltCon*list WeakTypeVar*list WeakCoerVar*list WeakExprVar*WeakExpr)) : ???(Tree
                       ??{ sac : _ & {scb : StrongCaseBranchWithVVs CoreVar CoreVarEqDecidable tc avars' sac &
-                        Expr (sac_gamma sac Γ) (sac_delta sac Γ avars' (weakCK'' Δ))(scbwv_xi scb ξ lev)(weakLT' (tbranches' @@  lev))}}) := 
+                        Expr (sac_gamma sac Γ) (sac_delta sac Γ avars' (weakCK'' Δ))(scbwv_xi scb ξ lev)(weakT' tbranches')(weakL' lev)}}) := 
                     match t with
                       | T_Leaf None           => OK []
                       | T_Leaf (Some (ac,extyvars,coervars,exprvars,ebranch)) => 
@@ -701,7 +699,7 @@ Definition weakExprToStrongExpr : forall
                     end) alts >>= fun tree =>
 
                     weakExprToStrongExpr Γ Δ φ ψ ξ ig (caseType tc avars') lev escrut >>= fun escrut' =>
-                      castExpr we "ECase" (τ@@lev) (ECase Γ Δ ξ lev tc tbranches' avars' escrut' tree)
+                      castExpr we "ECase" τ lev (ECase Γ Δ ξ lev tc tbranches' avars' escrut' tree)
     end)); try clear binds; try apply ConcatenableString.
   
     destruct case_some.
diff --git a/src/NaturalDeductionContext.v b/src/NaturalDeductionContext.v
new file mode 100644 (file)
index 0000000..57dfdb8
--- /dev/null
@@ -0,0 +1,255 @@
+(*********************************************************************************************************************************)
+(* NaturalDeductionContext:                                                                                                      *)
+(*                                                                                                                               *)
+(*   Manipulations of a context in natural deduction proofs.                                                                     *)
+(*                                                                                                                               *)
+(*********************************************************************************************************************************)
+
+Generalizable All Variables.
+Require Import Preamble.
+Require Import General.
+Require Import NaturalDeduction.
+
+Section NaturalDeductionContext.
+
+  (* Figure 3, production $\vdash_E$, Uniform rules *)
+  Inductive Arrange {T} : Tree ??T -> Tree ??T -> Type :=
+  | AId     : forall a        ,                Arrange           a                  a
+  | ACanL   : forall a        ,                Arrange  (    [],,a   )      (       a   )
+  | ACanR   : forall a        ,                Arrange  (    a,,[]   )      (       a   )
+  | AuCanL  : forall a        ,                Arrange  (       a    )      (  [],,a    )
+  | AuCanR  : forall a        ,                Arrange  (       a    )      (  a,,[]    )
+  | AAssoc  : forall a b c    ,                Arrange  (a,,(b,,c)   )      ((a,,b),,c  )
+  | AuAssoc  : forall a b c    ,                Arrange  ((a,,b),,c   )      ( a,,(b,,c) )
+  | AExch   : forall a b      ,                Arrange  (   (b,,a)   )      (  (a,,b)   )
+  | AWeak   : forall a        ,                Arrange  (       []   )      (       a   )
+  | ACont   : forall a        ,                Arrange  (  (a,,a)    )      (       a   )
+  | ALeft   : forall {h}{c} x , Arrange h c -> Arrange  (    x,,h    )      (       x,,c)
+  | ARight  : forall {h}{c} x , Arrange h c -> Arrange  (    h,,x    )      (       c,,x)
+  | AComp   : forall {a}{b}{c}, Arrange a b -> Arrange b c -> Arrange a c
+  .
+  
+  (* "Arrange" objects are parametric in the type of the leaves of the tree *)
+  Definition arrangeMap :
+    forall {T} (Σ₁ Σ₂:Tree ??T) {R} (f:T -> R),
+      Arrange Σ₁ Σ₂ ->
+      Arrange (mapOptionTree f Σ₁) (mapOptionTree f Σ₂).
+    intros.
+    induction X; simpl.
+    apply AId.
+    apply ACanL.
+    apply ACanR.
+    apply AuCanL.
+    apply AuCanR.
+    apply AAssoc.
+    apply AuAssoc.
+    apply AExch.
+    apply AWeak.
+    apply ACont.
+    apply ALeft; auto.
+    apply ARight; auto.
+    eapply AComp; [ apply IHX1 | apply IHX2 ].
+    Defined.
+  
+  (* a frequently-used Arrange - swap the middle two elements of a four-element sequence *)
+  Definition arrangeSwapMiddle {T} (a b c d:Tree ??T) :
+    Arrange ((a,,b),,(c,,d)) ((a,,c),,(b,,d)).
+    eapply AComp.
+    apply AuAssoc.
+    eapply AComp.
+    eapply ALeft.
+    eapply AComp.
+    eapply AAssoc.
+    eapply ARight.
+    apply AExch.
+    eapply AComp.
+    eapply ALeft.
+    eapply AuAssoc.
+    eapply AAssoc.
+    Defined.
+
+  (* like AExch, but works on nodes which are an Assoc away from being adjacent *)
+  Definition pivotContext {T} a b c : @Arrange T ((a,,b),,c) ((a,,c),,b) :=
+    AComp (AComp (AuAssoc _ _ _) (ALeft a (AExch c b))) (AAssoc _ _ _).
+
+  (* like AExch, but works on nodes which are an Assoc away from being adjacent *)  
+  Definition pivotContext' {T} a b c : @Arrange T (a,,(b,,c)) (b,,(a,,c)) :=
+    AComp (AComp (AAssoc _ _ _) (ARight c (AExch b a))) (AuAssoc _ _ _).
+  
+  Definition copyAndPivotContext {T} a b c : @Arrange T ((a,,b),,(c,,b)) ((a,,c),,b).
+    eapply AComp; [ idtac | apply (ALeft (a,,c) (ACont b)) ].
+    eapply AComp; [ idtac | apply AuAssoc ]. 
+    eapply AComp; [ idtac | apply (ARight b (pivotContext a b c)) ].
+    apply AAssoc.
+    Defined.
+
+  (* given any set of TreeFlags on a tree, we can Arrange all of the flagged nodes into the left subtree *)
+  Definition arrangePartition :
+    forall {T} (Σ:Tree ??T) (f:T -> bool),
+      Arrange Σ (dropT (mkFlags (liftBoolFunc false f) Σ),,( (dropT (mkFlags (liftBoolFunc false (bnot ○ f)) Σ)))).
+    intros.
+    induction Σ.
+      simpl.
+      destruct a.
+      simpl.
+      destruct (f t); simpl.
+      apply AuCanL.
+      apply AuCanR.
+      simpl.
+      apply AuCanL.
+      simpl in *.
+      eapply AComp; [ idtac | apply arrangeSwapMiddle ].
+      eapply AComp.
+      eapply ALeft.
+      apply IHΣ2.
+      eapply ARight.
+      apply IHΣ1.
+      Defined.
+
+  (* inverse of arrangePartition *)
+  Definition arrangeUnPartition :
+    forall {T} (Σ:Tree ??T) (f:T -> bool),
+      Arrange (dropT (mkFlags (liftBoolFunc false f) Σ),,( (dropT (mkFlags (liftBoolFunc false (bnot ○ f)) Σ)))) Σ.
+    intros.
+    induction Σ.
+      simpl.
+      destruct a.
+      simpl.
+      destruct (f t); simpl.
+      apply ACanL.
+      apply ACanR.
+      simpl.
+      apply ACanL.
+      simpl in *.
+      eapply AComp; [ apply arrangeSwapMiddle | idtac ].
+      eapply AComp.
+      eapply ALeft.
+      apply IHΣ2.
+      eapply ARight.
+      apply IHΣ1.
+      Defined.
+
+  (* we can decide if a tree consists exclusively of (T_Leaf None)'s *)
+  Definition decide_tree_empty : forall {T:Type}(t:Tree ??T),
+    sum { q:Tree unit & t = mapTree (fun _ => None) q } unit.
+    intro T.
+    refine (fix foo t :=
+      match t with
+        | T_Leaf x => _
+        | T_Branch b1 b2 => let b1' := foo b1 in let b2' := foo b2 in _
+      end).
+    intros.
+    destruct x.
+    right; apply tt.
+    left.
+      exists (T_Leaf tt).
+      auto.
+    destruct b1'.
+    destruct b2'.
+    destruct s.
+    destruct s0.
+    subst.
+    left.
+    exists (x,,x0).
+    reflexivity.
+    right; auto.
+    right; auto.
+    Defined.
+
+  (* if a tree is empty, we can Arrange it to [] *)
+  Definition arrangeCancelEmptyTree : forall {T}{A}(q:Tree A)(t:Tree ??T),
+    t = mapTree (fun _:A => None) q ->
+    Arrange t [].
+    intros T A q.
+    induction q; intros.
+      simpl in H.
+      rewrite H.
+      apply AId.
+    simpl in *.
+    destruct t; try destruct o; inversion H.
+      set (IHq1 _ H1) as x1.
+      set (IHq2 _ H2) as x2.
+      eapply AComp.
+        eapply ARight.
+        rewrite <- H1.
+        apply x1.
+      eapply AComp.
+        apply ACanL.
+        rewrite <- H2.
+        apply x2.
+      Defined.
+
+  (* if a tree is empty, we can Arrange it from [] *)
+  Definition arrangeUnCancelEmptyTree : forall {T}{A}(q:Tree A)(t:Tree ??T),
+    t = mapTree (fun _:A => None) q ->
+    Arrange [] t.
+    intros T A q.
+    induction q; intros.
+      simpl in H.
+      rewrite H.
+      apply AId.
+    simpl in *.
+    destruct t; try destruct o; inversion H.
+      set (IHq1 _ H1) as x1.
+      set (IHq2 _ H2) as x2.
+      eapply AComp.
+        apply AuCanL.
+      eapply AComp.
+        eapply ARight.
+        apply x1.
+      eapply AComp.
+        eapply ALeft.
+        apply x2.
+      rewrite H.
+      apply AId.
+      Defined.
+
+  (* given an Arrange from Σ₁ to Σ₂ and any predicate on tree nodes, we can construct an Arrange from (dropT Σ₁) to (dropT Σ₂) *)
+  Lemma arrangeDrop {T} pred
+    : forall (Σ₁ Σ₂: Tree ??T), Arrange Σ₁ Σ₂ -> Arrange (dropT (mkFlags pred Σ₁)) (dropT (mkFlags pred Σ₂)).
+
+    refine ((fix arrangeTake t1 t2 (arr:Arrange t1 t2) :=
+      match arr as R in Arrange A B return Arrange (dropT (mkFlags pred A)) (dropT (mkFlags pred B)) with
+        | AId  a               => let case_AId := tt    in AId _
+        | ACanL  a             => let case_ACanL := tt  in _
+        | ACanR  a             => let case_ACanR := tt  in _
+        | AuCanL a             => let case_AuCanL := tt in _
+        | AuCanR a             => let case_AuCanR := tt in _
+        | AAssoc a b c         => let case_AAssoc := tt in AAssoc _ _ _
+        | AuAssoc a b c         => let case_AuAssoc := tt in AuAssoc _ _ _
+        | AExch  a b           => let case_AExch := tt  in AExch _ _
+        | AWeak  a             => let case_AWeak := tt  in _
+        | ACont  a             => let case_ACont := tt  in _
+        | ALeft  a b c r'      => let case_ALeft := tt  in ALeft  _ (arrangeTake _ _ r')
+        | ARight a b c r'      => let case_ARight := tt in ARight _ (arrangeTake _ _ r')
+        | AComp  a b c r1 r2   => let case_AComp := tt  in AComp (arrangeTake _ _ r1) (arrangeTake _ _ r2)
+      end)); clear arrangeTake; intros.
+
+    destruct case_ACanL.
+      simpl; destruct (pred None); simpl; apply ACanL.
+
+    destruct case_ACanR.
+      simpl; destruct (pred None); simpl; apply ACanR.
+
+    destruct case_AuCanL.
+      simpl; destruct (pred None); simpl; apply AuCanL.
+
+    destruct case_AuCanR.
+      simpl; destruct (pred None); simpl; apply AuCanR.
+
+    destruct case_AWeak.
+      simpl; destruct (pred None); simpl; apply AWeak.
+
+    destruct case_ACont.
+      simpl; destruct (pred None); simpl; apply ACont.
+
+      Defined.
+
+  (* given an Arrange from Σ₁ to Σ₂ and any predicate on tree nodes, we can construct an Arrange from (takeT Σ₁) to (takeT Σ₂) *)
+  (*
+  Lemma arrangeTake {T} pred
+    : forall (Σ₁ Σ₂: Tree ??T), Arrange Σ₁ Σ₂ -> Arrange (takeT' (mkFlags pred Σ₁)) (takeT' (mkFlags pred Σ₂)).
+    unfold takeT'.
+    *)
+
+End NaturalDeductionContext.
index 5caf2dc..373519b 100644 (file)
--- a/src/PCF.v
+++ b/src/PCF.v
@@ -183,8 +183,8 @@ Section PCF.
       eapply nd_prod.
       apply nd_id.
       apply (PCF_Arrange [h] ([],,[h]) [h0]).
-      apply RuCanL.
-      eapply nd_comp; [ idtac | apply (PCF_Arrange ([],,a) a [h0]); apply RCanL ].
+      apply AuCanL.
+      eapply nd_comp; [ idtac | apply (PCF_Arrange ([],,a) a [h0]); apply ACanL ].
       apply nd_rule.
       (*
       set (@RLet Γ Δ [] (a@@@(ec::nil)) h0 h (ec::nil)) as q.
@@ -241,27 +241,27 @@ Section PCF.
   ; cnd_expand_right := fun a b c => PCF_right Γ Δ lev c a b }.
 
     intros; apply nd_rule. unfold PCFRule. simpl.
-      exists (RArrange _ _ _ _ _ (RCossa _ _ _)).
+      exists (RArrange _ _ _ _ _ (AuAssoc _ _ _)).
       apply (PCF_RArrange _ _ lev ((a,,b),,c) (a,,(b,,c)) x).
 
     intros; apply nd_rule. unfold PCFRule. simpl.
-      exists (RArrange _ _ _ _ _ (RAssoc _ _ _)).
+      exists (RArrange _ _ _ _ _ (AAssoc _ _ _)).
       apply (PCF_RArrange _ _ lev (a,,(b,,c)) ((a,,b),,c) x).
 
     intros; apply nd_rule. unfold PCFRule. simpl.
-      exists (RArrange _ _ _ _ _ (RCanL _)).
+      exists (RArrange _ _ _ _ _ (ACanL _)).
       apply (PCF_RArrange _ _ lev ([],,a) _ _).
 
     intros; apply nd_rule. unfold PCFRule. simpl.
-      exists (RArrange _ _ _ _ _ (RCanR _)).
+      exists (RArrange _ _ _ _ _ (ACanR _)).
       apply (PCF_RArrange _ _ lev (a,,[]) _ _).
 
     intros; apply nd_rule. unfold PCFRule. simpl.
-      exists (RArrange _ _ _ _ _ (RuCanL _)).
+      exists (RArrange _ _ _ _ _ (AuCanL _)).
       apply (PCF_RArrange _ _ lev _ ([],,a) _).
 
     intros; apply nd_rule. unfold PCFRule. simpl.
-      exists (RArrange _ _ _ _ _ (RuCanR _)).
+      exists (RArrange _ _ _ _ _ (AuCanR _)).
       apply (PCF_RArrange _ _ lev _ (a,,[]) _).
       Defined.