add crude support for flattening with LetRec and Case at level zero
[coq-hetmet.git] / src / HaskProofToStrong.v
index b14a53e..f50c586 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.
@@ -27,11 +28,11 @@ Section HaskProofToStrong.
 
   Definition judg2exprType (j:Judg) : Type :=
     match j with
-      (Γ > Δ > Σ |- τ) => forall (ξ:ExprVarResolver Γ) vars, Σ = mapOptionTree ξ vars ->
-        FreshM (ITree _ (fun t => Expr Γ Δ ξ t) τ)
+      (Γ > Δ > Σ |- τ @ l) => forall (ξ:ExprVarResolver Γ) vars, Σ = mapOptionTree ξ vars ->
+        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.
@@ -42,7 +43,7 @@ Section HaskProofToStrong.
     Defined.
 
   Lemma update_branches : forall Γ (ξ:VV -> LeveledHaskType Γ ★) lev l1 l2 q,
-    update_ξ ξ lev (app l1 l2) q = update_ξ (update_ξ ξ lev l2) lev l1 q.
+    update_xi ξ lev (app l1 l2) q = update_xi (update_xi ξ lev l2) lev l1 q.
     intros.
     induction l1.
       reflexivity.
@@ -122,7 +123,7 @@ Section HaskProofToStrong.
   Lemma fresh_lemma'' Γ 
     : forall types ξ lev, 
     FreshM { varstypes : _
-      |  mapOptionTree (update_ξ(Γ:=Γ)   ξ lev (leaves varstypes)) (mapOptionTree (@fst _ _) varstypes) = (types @@@ lev)
+      |  mapOptionTree (update_xi(Γ:=Γ)   ξ lev (leaves varstypes)) (mapOptionTree (@fst _ _) varstypes) = (types @@@ lev)
       /\ distinct (leaves (mapOptionTree (@fst _ _) varstypes)) }.
   admit.
   Defined.
@@ -130,8 +131,8 @@ Section HaskProofToStrong.
   Lemma fresh_lemma' Γ 
     : forall types vars Σ ξ lev, Σ = mapOptionTree ξ vars ->
     FreshM { varstypes : _
-      |  mapOptionTree (update_ξ(Γ:=Γ) ξ lev (leaves varstypes)) vars = Σ
-      /\ mapOptionTree (update_ξ       ξ lev (leaves varstypes)) (mapOptionTree (@fst _ _) varstypes) = (types @@@ lev)
+      |  mapOptionTree (update_xi(Γ:=Γ) ξ lev (leaves varstypes)) vars = Σ
+      /\ mapOptionTree (update_xi       ξ lev (leaves varstypes)) (mapOptionTree (@fst _ _) varstypes) = (types @@@ lev)
       /\ distinct (leaves (mapOptionTree (@fst _ _) varstypes)) }.
     induction types.
       intros; destruct a.
@@ -164,7 +165,7 @@ Section HaskProofToStrong.
         intros vars Σ ξ lev pf; refine (bind x2 = IHtypes2 vars Σ ξ lev pf; _).
           apply FreshMon.
           destruct x2 as [vt2 [pf21 [pf22 pfdist]]].
-          refine (bind x1 = IHtypes1 (vars,,(mapOptionTree (@fst _ _) vt2)) (Σ,,(types2@@@lev)) (update_ξ ξ lev
+          refine (bind x1 = IHtypes1 (vars,,(mapOptionTree (@fst _ _) vt2)) (Σ,,(types2@@@lev)) (update_xi ξ lev
             (leaves vt2)) _ _; return _).
           apply FreshMon.
           simpl.
@@ -204,8 +205,8 @@ Section HaskProofToStrong.
   Lemma fresh_lemma Γ ξ vars Σ Σ' lev
     : Σ = mapOptionTree ξ vars ->
     FreshM { vars' : _
-      |  mapOptionTree (update_ξ(Γ:=Γ) ξ lev ((vars',Σ')::nil)) vars = Σ
-      /\ mapOptionTree (update_ξ ξ lev ((vars',Σ')::nil)) [vars'] = [Σ' @@ lev] }.
+      |  mapOptionTree (update_xi(Γ:=Γ) ξ lev ((vars',Σ')::nil)) vars = Σ
+      /\ mapOptionTree (update_xi ξ lev ((vars',Σ')::nil)) [vars'] = [Σ' @@ lev] }.
     intros.
     set (fresh_lemma' Γ [Σ'] vars Σ ξ lev H) as q.
     refine (q >>>= fun q' => return _).
@@ -222,60 +223,64 @@ Section HaskProofToStrong.
       inversion pf2.
     Defined.
 
-  Definition ujudg2exprType Γ (ξ:ExprVarResolver Γ)(Δ:CoercionEnv Γ) Σ τ : Type :=
-    forall vars, Σ = mapOptionTree ξ vars -> FreshM (ITree _ (fun t => Expr Γ Δ ξ t) τ).
+  Definition ujudg2exprType Γ (ξ:ExprVarResolver Γ)(Δ:CoercionEnv Γ) Σ τ l : Type :=
+    forall vars, Σ = mapOptionTree ξ vars -> FreshM (ITree _ (fun t => Expr Γ Δ ξ t l) τ).
 
-  Definition urule2expr  : forall Γ Δ h j t (r:@Arrange _ h j) (ξ:VV -> LeveledHaskType Γ ★),
-    ujudg2exprType Γ ξ Δ h t ->
-    ujudg2exprType Γ ξ Δ j t
+  Definition urule2expr  : forall Γ Δ h j t l (r:@Arrange _ h j) (ξ:VV -> LeveledHaskType Γ ★),
+    ujudg2exprType Γ ξ Δ h t l ->
+    ujudg2exprType Γ ξ Δ j t l
     .
     intros Γ Δ.
-      refine (fix urule2expr h j t (r:@Arrange _ h j) ξ {struct r} : 
-    ujudg2exprType Γ ξ Δ h t ->
-    ujudg2exprType Γ ξ Δ j t :=
+      refine (fix urule2expr h j t l (r:@Arrange _ h j) ξ {struct r} : 
+    ujudg2exprType Γ ξ Δ h t l ->
+    ujudg2exprType Γ ξ Δ j t l :=
         match r as R in Arrange H C return
-    ujudg2exprType Γ ξ Δ H t ->
-    ujudg2exprType Γ ξ Δ C t
+    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)
-          | 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_RCanL.
+      destruct case_AId.
+        apply X.
+
+      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.
@@ -283,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.
@@ -291,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)).
@@ -312,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).
@@ -327,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).
@@ -342,16 +347,16 @@ Section HaskProofToStrong.
           simpl.
           reflexivity.
 
-      destruct case_RComp.
+      destruct case_AComp.
         apply e2.
         apply e1.
         apply X.
         Defined.
 
   Definition letrec_helper Γ Δ l (varstypes:Tree ??(VV * HaskType Γ ★)) ξ' :
-    ITree (LeveledHaskType Γ ★)
-         (fun t : LeveledHaskType Γ ★ => Expr Γ Δ ξ' t)
-         (mapOptionTree (ξ' ○ (@fst _ _)) varstypes)
+    ITree (HaskType Γ ★)
+         (fun t : HaskType Γ ★ => Expr Γ Δ ξ' t l)
+         (mapOptionTree (unlev ○ ξ' ○ (@fst _ _)) varstypes)
          -> ELetRecBindings Γ Δ ξ' l varstypes.
     intros.
     induction varstypes.
@@ -367,6 +372,8 @@ Section HaskProofToStrong.
       simpl.
       destruct (eqd_dec h0 l).
         rewrite <- e0.
+        simpl in X.
+        subst.
         apply X.
       apply (Prelude_error "level mismatch; should never happen").
       apply (Prelude_error "letrec type mismatch; should never happen").
@@ -386,9 +393,9 @@ Section HaskProofToStrong.
     exists x; auto.
     Defined.
 
-  Definition fix_indexing X (F:X->Type)(J:X->Type)(t:Tree ??{ x:X & F x })
-    :  ITree { x:X & F x } (fun x => J (projT1 x))                                t
-    -> ITree X             (fun x:X => J x)   (mapOptionTree (@projT1 _ _) t).
+  Definition fix_indexing X Y (J:X->Type)(t:Tree ??(X*Y))
+    :  ITree (X * Y) (fun x => J (fst x))                                t
+    -> ITree X       (fun x:X => J x)   (mapOptionTree (@fst _ _) t).
     intro it.
     induction it; simpl in *.
     apply INone.
@@ -411,11 +418,13 @@ Section HaskProofToStrong.
     Defined.
   
   Definition case_helper tc Γ Δ lev tbranches avars ξ :
-    forall pcb:{sac : StrongAltCon & ProofCaseBranch tc Γ Δ lev tbranches avars sac},
-     prod (judg2exprType (pcb_judg (projT2 pcb))) {vars' : Tree ??VV & pcb_freevars (projT2 pcb) = mapOptionTree ξ vars'} ->
+    forall pcb:(StrongAltCon * Tree ??(LeveledHaskType Γ ★)),
+     prod (judg2exprType (@pcb_judg tc Γ Δ lev tbranches avars (fst pcb) (snd pcb)))
+     {vars' : Tree ??VV & (snd pcb) = mapOptionTree ξ vars'} ->
      ((fun sac => FreshM
        { scb : StrongCaseBranchWithVVs VV eqdec_vv tc avars sac
-         & Expr (sac_Γ sac Γ) (sac_Δ sac Γ avars (weakCK'' Δ)) (scbwv_ξ scb ξ lev) (weakLT' (tbranches @@  lev)) }) (projT1 pcb)).
+         & Expr (sac_gamma sac Γ) (sac_delta sac Γ avars (weakCK'' Δ)) (scbwv_xi scb ξ lev)
+         (weakT' tbranches) (weakL' lev) }) (fst pcb)).
     intro pcb.
     intro X.
     simpl in X.
@@ -427,7 +436,7 @@ Section HaskProofToStrong.
     destruct s as [vars vars_pf].
 
     refine (bind localvars = fresh_lemma' _ (unleaves  (vec2list (sac_types sac _ avars))) vars 
-      (mapOptionTree weakLT' (pcb_freevars pcb)) (weakLT' ○ ξ) (weakL' lev) _  ; _).
+      (mapOptionTree weakLT' pcb) (weakLT' ○ ξ) (weakL' lev) _  ; _).
       apply FreshMon.
       rewrite vars_pf.
       rewrite <- mapOptionTree_compose.
@@ -441,10 +450,10 @@ Section HaskProofToStrong.
     cut (distinct (vec2list localvars'')). intro H'''.
     set (@Build_StrongCaseBranchWithVVs _ _ _ _ avars sac localvars'' H''') as scb.
 
-    refine (bind q = (f (scbwv_ξ scb ξ lev) (vars,,(unleaves (vec2list (scbwv_exprvars scb)))) _) ; return _).
+    refine (bind q = (f (scbwv_xi scb ξ lev) (vars,,(unleaves (vec2list (scbwv_exprvars scb)))) _) ; return _).
       apply FreshMon.
       simpl.
-      unfold scbwv_ξ.
+      unfold scbwv_xi.
       rewrite vars_pf.
       rewrite <- mapOptionTree_compose.
       clear localvars_pf1.
@@ -462,14 +471,15 @@ Section HaskProofToStrong.
     Defined.
 
   Definition gather_branch_variables
-    Γ Δ (ξ:VV -> LeveledHaskType Γ ★) tc avars tbranches lev (alts:Tree ?? {sac : StrongAltCon &
-                ProofCaseBranch tc Γ Δ lev tbranches avars sac})
+    Γ Δ
+    (ξ:VV -> LeveledHaskType Γ ★) tc avars tbranches lev
+    (alts:Tree ??(@StrongAltCon tc * Tree ??(LeveledHaskType Γ ★)))
     :
     forall vars,
-    mapOptionTreeAndFlatten (fun x => pcb_freevars(Γ:=Γ) (projT2 x)) alts = mapOptionTree ξ vars
-    -> ITree Judg judg2exprType (mapOptionTree (fun x => pcb_judg (projT2 x)) alts)
-    -> ITree _ (fun q => prod (judg2exprType (pcb_judg (projT2 q))) 
-      { vars' : _ & pcb_freevars (projT2 q) = mapOptionTree ξ vars' })
+    mapOptionTreeAndFlatten (fun x => snd x) alts = mapOptionTree ξ vars
+    -> ITree Judg judg2exprType (mapOptionTree (fun x => @pcb_judg tc Γ Δ lev avars tbranches (fst x) (snd x)) alts)
+    -> ITree _ (fun q => prod (judg2exprType (@pcb_judg tc Γ Δ lev avars tbranches (fst q) (snd q))) 
+      { vars' : _ & (snd q) = mapOptionTree ξ vars' })
   alts.
     induction alts;
     intro vars;
@@ -479,7 +489,7 @@ Section HaskProofToStrong.
     simpl in *.
     apply ileaf in source.
     apply ILeaf.
-    destruct s as [sac pcb].
+    destruct p as [sac pcb].
     simpl in *.
     split.
     intros.
@@ -501,13 +511,249 @@ 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.
 
     intros h j r.
 
       refine (match r as R in Rule H C return ITree _ judg2exprType H -> ITree _ judg2exprType C with
-      | RArrange a b c d e  r         => let case_RURule := tt        in _
+      | RArrange a b c d e l r        => let case_RURule := tt        in _
       | RNote   Γ Δ Σ τ l n           => let case_RNote := tt         in _
       | RLit    Γ Δ l     _           => let case_RLit := tt          in _
       | RVar    Γ Δ σ         p       => let case_RVar := tt          in _
@@ -519,9 +765,10 @@ Section HaskProofToStrong.
       | RAppCo  Γ Δ Σ κ σ₁ σ₂ γ σ l   => let case_RAppCo := tt        in _
       | RAbsCo  Γ Δ Σ κ σ  σ₁ σ₂  y   => let case_RAbsCo := tt        in _
       | RApp    Γ Δ Σ₁ Σ₂ tx te p     => let case_RApp := tt          in _
-      | RLet    Γ Δ Σ₁ Σ₂ σ₁ σ₂ p     => let case_RLet := tt          in _
-      | RJoin Γ p lri m x q   => let case_RJoin := tt in _
-      | RVoid _ _               => let case_RVoid := tt   in _
+      | RCut    Γ Δ Σ Σ₁ Σ₁₂ Σ₂ Σ₃ l  => let case_RCut := tt          in _
+      | RLeft   Γ Δ Σ₁ Σ₂  Σ     l    => let case_RLeft := tt in _
+      | RRight  Γ Δ Σ₁ Σ₂  Σ     l    => let case_RRight := tt in _
+      | RVoid   _ _ l                 => let case_RVoid := tt   in _
       | RBrak   Σ a b c n m           => let case_RBrak := tt         in _
       | REsc    Σ a b c n m           => let case_REsc := tt          in _
       | RCase   Γ Δ lev tc Σ avars tbranches alts => let case_RCase := tt         in _
@@ -530,12 +777,10 @@ Section HaskProofToStrong.
 
     destruct case_RURule.
       apply ILeaf. simpl. intros.
-      set (@urule2expr a b _ _ e r0 ξ) as q.
-      set (fun z => q z) as q'.
-      simpl in q'.
-      apply q' with (vars:=vars).
-      clear q' q.
+      set (@urule2expr a b _ _ e l r0 ξ) as q.
       unfold ujudg2exprType.
+      unfold ujudg2exprType in q.
+      apply q with (vars:=vars).
       intros.
       apply X_ with (vars:=vars0).
       auto.
@@ -562,8 +807,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.
@@ -576,7 +824,7 @@ Section HaskProofToStrong.
     refine (fresh_lemma _ ξ vars _ tx x H >>>= (fun pf => _)).
     apply FreshMon.
     destruct pf as [ vnew [ pf1 pf2 ]].
-    set (update_ξ ξ x (((vnew, tx  )) :: nil)) as ξ' in *.
+    set (update_xi ξ x (((vnew, tx  )) :: nil)) as ξ' in *.
     refine (X_ ξ' (vars,,[vnew]) _ >>>= _).
     apply FreshMon.
     simpl.
@@ -600,19 +848,6 @@ Section HaskProofToStrong.
     apply ileaf in X. simpl in X.
     apply X.
 
-  destruct case_RJoin.
-    apply ILeaf; simpl; intros.
-    inversion X_.
-    apply ileaf in X.
-    apply ileaf in X0.
-    simpl in *.
-    destruct vars; inversion H.
-    destruct o; inversion H3.
-    refine (X ξ vars1 H3 >>>= fun X' => X0 ξ vars2 H4 >>>= fun X0' => return _).
-    apply FreshMon.
-    apply FreshMon.
-    apply IBranch; auto.
-
   destruct case_RApp.    
     apply ILeaf.
     inversion X_.
@@ -634,39 +869,65 @@ Section HaskProofToStrong.
     simpl in *.
     apply (EApp _ _ _ _ _ _ q1' q2').
 
-  destruct case_RLet.
-    apply ILeaf.
-    simpl in *; intros.
-    destruct vars; try destruct o; inversion H.
-    refine (fresh_lemma _ ξ vars1 _ σ₂ p H1 >>>= (fun pf => _)).
-    apply FreshMon.
-    destruct pf as [ vnew [ pf1 pf2 ]].
-    set (update_ξ ξ p (((vnew, σ₂ )) :: nil)) as ξ' in *.
+  destruct case_RCut.
+    apply rassoc.
+    apply swapr.
+    apply rassoc'.
+
     inversion X_.
-    apply ileaf in X.
-    apply ileaf in X0.
+    subst.
+    clear X_.
+
+    apply rassoc' in X0.
+    apply swapr in X0.
+    apply rassoc in X0.
+
+    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 *.
-    refine (X ξ  vars2 _ >>>= fun X0' => _).
+    apply ILeaf.
+    simpl.
+    intros.
+    refine (q1 _ _ H >>>= fun q1' => q2 _ _ H >>>= fun q2' => return _).
     apply FreshMon.
-    auto.
+    apply FreshMon.
+    apply IBranch; auto.
 
-    refine (X0 ξ' (vars1,,[vnew]) _ >>>= fun X1' => _).
+  destruct case_RLeft.
+    apply ILeaf.
+    simpl; intros.
+    destruct vars; try destruct o; inversion H.
+    refine (X_ _ _ H2 >>>= fun X' => return _).
     apply FreshMon.
-    rewrite H1.
-    simpl.
-    rewrite pf2.
-    rewrite pf1.
-    rewrite H1.
-    reflexivity.
+    apply IBranch.
+    eapply vartree.
+    apply H1.
+    apply X'.
 
-    refine (return _).
+  destruct case_RRight.
     apply ILeaf.
-    apply ileaf in X0'.
-    apply ileaf in X1'.
-    simpl in *.
-    apply ELet with (ev:=vnew)(tv:=σ₂).
-    apply X0'.
-    apply X1'.
+    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.
@@ -704,8 +965,8 @@ Section HaskProofToStrong.
     apply ILeaf; simpl; intros.
     refine (bind ξvars = fresh_lemma' _ y _ _ _ t H; _). apply FreshMon.
     destruct ξvars as [ varstypes [ pf1[ pf2 pfdist]]].
-    refine (X_ ((update_ξ ξ t (leaves varstypes)))
-      (vars,,(mapOptionTree (@fst _ _) varstypes)) _ >>>= fun X => return _); clear X_.  apply FreshMon.
+    refine (X_ ((update_xi ξ t (leaves varstypes)))
+      ((mapOptionTree (@fst _ _) varstypes),,vars) _ >>>= fun X => return _); clear X_.  apply FreshMon.
     simpl.
     rewrite pf2.
     rewrite pf1.
@@ -716,11 +977,19 @@ Section HaskProofToStrong.
     apply (@ELetRec _ _ _ _ _ _ _ varstypes).
     auto.
     apply (@letrec_helper Γ Δ t varstypes).
-    rewrite <- pf2 in X1.
     rewrite mapOptionTree_compose.
+    rewrite mapOptionTree_compose.
+    rewrite pf2.
+    replace ((mapOptionTree unlev (y @@@ t))) with y.
+      apply X0.
+      clear pf1 X0 X1 pfdist pf2 vars varstypes.
+      induction y; try destruct a; auto.
+      rewrite IHy1 at 1.
+      rewrite IHy2 at 1.
+      reflexivity.
+    apply ileaf in X1.
+    simpl in X1.
     apply X1.
-    apply ileaf in X0.
-    apply X0.
 
   destruct case_RCase.
     apply ILeaf; simpl; intros.
@@ -765,41 +1034,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_ξ ξ 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 [] [Γ > Δ > Σ |- [τ]] ->
-    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.
@@ -815,7 +1052,9 @@ Section HaskProofToStrong.
     auto.
     refine (return OK _).
     exists ξ.
-    apply (ileaf it).
+    apply ileaf in it.
+    simpl in it.
+    apply it.
     apply INone.
     Defined.