From b750ea4095c9c2195a056c9987bac03b5d23c01d Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 10 Jan 2025 16:51:49 -0500
Subject: [PATCH 01/13] Initial commit (still tweaking the rules)

---
 syntax.sig                   |  4 ++
 theories/Autosubst2/syntax.v | 86 +++++++++++++++++++++++++++++++++++-
 theories/fp_red.v            | 21 ++++++++-
 3 files changed, 109 insertions(+), 2 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index 4549f7f..d2354e2 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -2,6 +2,7 @@ nat  : Type
 Tm(VarTm) : Type
 PTag : Type
 TTag : Type
+bool : Type
 
 TPi : TTag
 TSig : TTag
@@ -14,3 +15,6 @@ Proj : PTag -> Tm -> Tm
 TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm
 Bot : Tm
 Univ : nat -> Tm
+Bool : Tm
+BVal : bool -> Tm
+If : Tm -> Tm -> Tm -> Tm
\ No newline at end of file
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index b972476..ffb7c7c 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -41,7 +41,10 @@ Inductive Tm (n_Tm : nat) : Type :=
   | Proj : PTag -> Tm n_Tm -> Tm n_Tm
   | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
   | Bot : Tm n_Tm
-  | Univ : nat -> Tm n_Tm.
+  | Univ : nat -> Tm n_Tm
+  | Bool : Tm n_Tm
+  | BVal : bool -> Tm n_Tm
+  | If : Tm n_Tm -> Tm n_Tm -> Tm n_Tm -> Tm n_Tm.
 
 Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)}
   (H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0.
@@ -94,6 +97,27 @@ Proof.
 exact (eq_trans eq_refl (ap (fun x => Univ m_Tm x) H0)).
 Qed.
 
+Lemma congr_Bool {m_Tm : nat} : Bool m_Tm = Bool m_Tm.
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma congr_BVal {m_Tm : nat} {s0 : bool} {t0 : bool} (H0 : s0 = t0) :
+  BVal m_Tm s0 = BVal m_Tm t0.
+Proof.
+exact (eq_trans eq_refl (ap (fun x => BVal m_Tm x) H0)).
+Qed.
+
+Lemma congr_If {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {s2 : Tm m_Tm}
+  {t0 : Tm m_Tm} {t1 : Tm m_Tm} {t2 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1)
+  (H2 : s2 = t2) : If m_Tm s0 s1 s2 = If m_Tm t0 t1 t2.
+Proof.
+exact (eq_trans
+         (eq_trans (eq_trans eq_refl (ap (fun x => If m_Tm x s1 s2) H0))
+            (ap (fun x => If m_Tm t0 x s2) H1))
+         (ap (fun x => If m_Tm t0 t1 x) H2)).
+Qed.
+
 Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) :
   fin (S m) -> fin (S n).
 Proof.
@@ -118,6 +142,10 @@ Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
       TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
   | Bot _ => Bot n_Tm
   | Univ _ s0 => Univ n_Tm s0
+  | Bool _ => Bool n_Tm
+  | BVal _ s0 => BVal n_Tm s0
+  | If _ s0 s1 s2 =>
+      If n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) (ren_Tm xi_Tm s2)
   end.
 
 Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) :
@@ -145,6 +173,11 @@ Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
       TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
   | Bot _ => Bot n_Tm
   | Univ _ s0 => Univ n_Tm s0
+  | Bool _ => Bool n_Tm
+  | BVal _ s0 => BVal n_Tm s0
+  | If _ s0 s1 s2 =>
+      If n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
+        (subst_Tm sigma_Tm s2)
   end.
 
 Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm)
@@ -185,6 +218,11 @@ subst_Tm sigma_Tm s = s :=
         (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
+        (idSubst_Tm sigma_Tm Eq_Tm s2)
   end.
 
 Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n)
@@ -229,6 +267,11 @@ Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
            (upExtRen_Tm_Tm _ _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0)
+        (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s2)
   end.
 
 Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm)
@@ -274,6 +317,11 @@ Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
            s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (ext_Tm sigma_Tm tau_Tm Eq_Tm s0)
+        (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) (ext_Tm sigma_Tm tau_Tm Eq_Tm s2)
   end.
 
 Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
@@ -319,6 +367,12 @@ Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
            (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0)
+        (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
+        (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s2)
   end.
 
 Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat}
@@ -375,6 +429,12 @@ Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
            (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0)
+        (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
+        (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s2)
   end.
 
 Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@@ -452,6 +512,12 @@ ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
            (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0)
+        (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
+        (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s2)
   end.
 
 Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@@ -530,6 +596,12 @@ subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
            (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0)
+        (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
+        (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s2)
   end.
 
 Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
@@ -646,6 +718,12 @@ Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
            (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2)
   | Bot _ => congr_Bot
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bool _ => congr_Bool
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | If _ s0 s1 s2 =>
+      congr_If (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0)
+        (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
+        (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s2)
   end.
 
 Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
@@ -844,6 +922,12 @@ Core.
 
 Arguments VarTm {n_Tm}.
 
+Arguments If {n_Tm}.
+
+Arguments BVal {n_Tm}.
+
+Arguments Bool {n_Tm}.
+
 Arguments Univ {n_Tm}.
 
 Arguments Bot {n_Tm}.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index d2fbdf6..9424da6 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -33,6 +33,8 @@ Module Par.
     R b0 b1 ->
     R c0 c1 ->
     R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1))
+  | AppBool b a :
+    R (App (BVal b) a) (BVal b)
   | ProjAbs p a0 a1 :
     R a0 a1 ->
     R (Proj p (Abs a0)) (Abs (Proj p a1))
@@ -40,7 +42,24 @@ Module Par.
     R a0 a1 ->
     R b0 b1 ->
     R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
-
+  | ProjBool p b :
+    R (Proj p (BVal b)) (BVal b)
+  | IfAbs a0 a1 b0 b1 c0 c1 d0 d1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R d0 d1 ->
+    R (If (Abs a0) b0 c0) (Abs (If a0 (ren_Tm shift b0) (ren_Tm shift c0)))
+  | IfPair a0 a1 b0 b1 c0 c1 d0 d1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R d0 d1 ->
+    R (If (Pair a0 b0) c0 d0) (Pair (If a0 c0 d0) (If b0 c0 d0))
+  | IfBool a b0 b1 c0 c1 :
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (BVal a) b0 c0) (if a then b1 else c1)
   (****************** Eta ***********************)
   | AppEta a0 a1 :
     R a0 a1 ->

From bec803949f4da3522031a185cdfcd522be675b82 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 10 Jan 2025 20:51:59 -0500
Subject: [PATCH 02/13] Finish Par and RPar lemmas

---
 theories/fp_red.v | 117 +++++++++++++++++++++++++++++++++++++---------
 1 file changed, 96 insertions(+), 21 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 9424da6..c8076e1 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -33,8 +33,8 @@ Module Par.
     R b0 b1 ->
     R c0 c1 ->
     R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1))
-  | AppBool b a :
-    R (App (BVal b) a) (BVal b)
+  (* | AppBool b a : *)
+  (*   R (App (BVal b) a) (BVal b) *)
   | ProjAbs p a0 a1 :
     R a0 a1 ->
     R (Proj p (Abs a0)) (Abs (Proj p a1))
@@ -42,20 +42,20 @@ Module Par.
     R a0 a1 ->
     R b0 b1 ->
     R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
-  | ProjBool p b :
-    R (Proj p (BVal b)) (BVal b)
-  | IfAbs a0 a1 b0 b1 c0 c1 d0 d1 :
-    R a0 a1 ->
-    R b0 b1 ->
-    R c0 c1 ->
-    R d0 d1 ->
-    R (If (Abs a0) b0 c0) (Abs (If a0 (ren_Tm shift b0) (ren_Tm shift c0)))
-  | IfPair a0 a1 b0 b1 c0 c1 d0 d1 :
-    R a0 a1 ->
-    R b0 b1 ->
-    R c0 c1 ->
-    R d0 d1 ->
-    R (If (Pair a0 b0) c0 d0) (Pair (If a0 c0 d0) (If b0 c0 d0))
+  (* | ProjBool p b : *)
+  (*   R (Proj p (BVal b)) (BVal b) *)
+  (* | IfAbs a0 a1 b0 b1 c0 c1 d0 d1 : *)
+  (*   R a0 a1 -> *)
+  (*   R b0 b1 -> *)
+  (*   R c0 c1 -> *)
+  (*   R d0 d1 -> *)
+  (*   R (If (Abs a0) b0 c0) (Abs (If a0 (ren_Tm shift b0) (ren_Tm shift c0))) *)
+  (* | IfPair a0 a1 b0 b1 c0 c1 d0 d1 : *)
+  (*   R a0 a1 -> *)
+  (*   R b0 b1 -> *)
+  (*   R c0 c1 -> *)
+  (*   R d0 d1 -> *)
+  (*   R (If (Pair a0 b0) c0 d0) (Pair (If a0 c0 d0) (If b0 c0 d0)) *)
   | IfBool a b0 b1 c0 c1 :
     R b0 b1 ->
     R c0 c1 ->
@@ -92,7 +92,16 @@ Module Par.
   | BotCong :
     R Bot Bot
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BoolCong :
+    R Bool Bool
+  | BValCong b :
+    R (BVal b) (BVal b)
+  | IfCong a0 a1 b0 b1 c0 c1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If a0 b0 c0) (If a1 b1 c1).
 
   Lemma refl n (a : Tm n) : R a a.
     elim : n /a; hauto ctrs:R.
@@ -112,6 +121,13 @@ Module Par.
     R (Proj p (Pair a0 b0)) t.
   Proof.  move => > ->. apply ProjPair. Qed.
 
+  Lemma IfBool' n a b0 b1 c0 c1 u :
+    u = (if a then (b1 : Tm n) else c1) ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (BVal a) b0 c0) u.
+  Proof. move => ->. apply IfBool. Qed.
+
   Lemma AppEta' n (a0 a1 b : Tm n) :
     b = (Abs (App (ren_Tm shift a1) (VarTm var_zero))) ->
     R a0 a1 ->
@@ -126,7 +142,7 @@ Module Par.
     move => *; apply : AppAbs'; eauto; by asimpl.
     all : match goal with
           | [ |- context[var_zero]] => move => *; apply : AppEta'; eauto; by asimpl
-          | _ => qauto ctrs:R use:ProjPair'
+          | _ => qauto ctrs:R use:ProjPair', IfBool'
           end.
   Qed.
 
@@ -143,6 +159,7 @@ Module Par.
     - hauto lq:on rew:off ctrs:R.
     - hauto l:on inv:option use:renaming ctrs:R.
     - hauto lq:on use:ProjPair'.
+    - hauto q:on use:IfBool'.
     - move => n a0 a1 ha iha m ρ0 ρ1 hρ /=.
       apply : AppEta'; eauto. by asimpl.
     - hauto lq:on ctrs:R.
@@ -154,6 +171,9 @@ Module Par.
     - hauto l:on inv:option ctrs:R use:renaming.
     - sfirstorder.
     - sfirstorder.
+    - sfirstorder.
+    - sfirstorder.
+    - hauto lq:on ctrs:R.
   Qed.
 
   Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
@@ -193,6 +213,12 @@ Module Par.
       move : ihb => [c0 [? ?]]. subst.
       eexists. split. by eauto using ProjPair.
       hauto q:on.
+    - move => n a b0 b1 c0 c1 hb ihb hc ihc m ξ []//= p0 p1 p2 [h *]. subst.
+      spec_refl.
+      move : ihb => [b0 [? ?]].
+      move : ihc => [c0 [? ?]]. subst.
+      case : p0 h => //=.
+      hauto q:on use:IfBool'.
     - move => n a0 a1 ha iha m ξ a ?. subst.
       spec_refl. move : iha => [a0 [? ?]]. subst.
       eexists. split. apply AppEta; eauto.
@@ -233,6 +259,11 @@ Module Par.
       by asimpl.
     - move => n n0 ξ []//=. hauto l:on.
     - move => n i n0 ξ []//=. hauto l:on.
+    - move => n m ξ []//=. hauto l:on.
+    - move => n b m ξ []//=. hauto l:on.
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ξ []//= t t0 t1[*]. subst.
+      spec_refl.
+      qauto l:on ctrs:R.
   Qed.
 
 End Par.
@@ -321,6 +352,10 @@ Module RPar.
     R a0 a1 ->
     R b0 b1 ->
     R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
+  | IfBool a b0 b1 c0 c1 :
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (BVal a) b0 c0) (if a then b1 else c1)
 
 
   (*************** Congruence ********************)
@@ -346,7 +381,16 @@ Module RPar.
   | BotCong :
     R Bot Bot
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BoolCong :
+    R Bool Bool
+  | BValCong b :
+    R (BVal b) (BVal b)
+  | IfCong a0 a1 b0 b1 c0 c1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If a0 b0 c0) (If a1 b1 c1).
 
   Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
 
@@ -369,13 +413,21 @@ Module RPar.
     R (Proj p (Pair a0 b0)) t.
   Proof.  move => > ->. apply ProjPair. Qed.
 
+  Lemma IfBool' n a b0 b1 c0 c1 u :
+    u = (if a then (b1 : Tm n) else c1) ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (BVal a) b0 c0) u.
+  Proof. move => ->. apply IfBool. Qed.
+
+
   Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
     R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
     move => *; apply : AppAbs'; eauto; by asimpl.
-    all : qauto ctrs:R use:ProjPair'.
+    all : qauto ctrs:R use:ProjPair', IfBool'.
   Qed.
 
   Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> Tm m) (ξ : fin m -> fin p) :
@@ -406,6 +458,7 @@ Module RPar.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R use:ProjPair' use:morphing_up.
+    - hauto lq:on ctrs:R use:IfBool' use:morphing_up.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R use:morphing_up.
@@ -414,6 +467,9 @@ Module RPar.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
   Qed.
 
   Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
@@ -434,7 +490,7 @@ Module RPar.
     var_or_bot a ->
     a = b -> ~~ var_or_bot b -> False.
   Proof.
-    hauto lq:on inv:Tm.
+    hauto l:on inv:Tm.
   Qed.
 
   Lemma var_or_bot_up n m (ρ : fin n -> Tm m) :
@@ -504,6 +560,14 @@ Module RPar.
       move : ihb => [c0 [? ?]]. subst.
       eexists. split. by eauto using ProjPair.
       hauto q:on.
+    - move => n a b0 b1 c0 c1 hb ihb hc ihc m ρ hρ []//=; first by antiimp.
+      move => t t0 t1 [h *]. subst.
+      case : t h => //=; first by antiimp.
+      move => b [*]. subst.
+      have {}/ihb := (hρ) => ihb.
+      have {}/ihc := (hρ) => ihc.
+      spec_refl.
+      hauto q:on use:IfBool.
     - move => n i m ρ hρ []//=.
       hauto l:on.
     - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp.
@@ -553,6 +617,17 @@ Module RPar.
     - hauto q:on ctrs:R inv:Tm.
     - move => n i n0 ρ hρ []//=; first by antiimp.
       hauto l:on.
+    - move => n m ρ hρ []//=; first by antiimp.
+      hauto lq:on ctrs:R.
+    - move => n b m ρ hρ []//; first by antiimp.
+      hauto lq:on ctrs:R.
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ hρ []//=; first by antiimp.
+      move => t t0 t1 [*]. subst.
+      have {}/iha := (hρ) => iha.
+      have {}/ihb := (hρ) => ihb.
+      have {}/ihc := (hρ) => ihc.
+      spec_refl.
+      hauto lq:on ctrs:R.
   Qed.
 End RPar.
 

From 53bef034ad74a93ff6b98e330b48c4139e4f3962 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 10 Jan 2025 22:28:53 -0500
Subject: [PATCH 03/13] Add the "junk" rules for boolean

---
 theories/fp_red.v | 897 ++++++++++++++++++++++++++--------------------
 1 file changed, 509 insertions(+), 388 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index c8076e1..a8c8a3a 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -33,8 +33,8 @@ Module Par.
     R b0 b1 ->
     R c0 c1 ->
     R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1))
-  (* | AppBool b a : *)
-  (*   R (App (BVal b) a) (BVal b) *)
+  | AppBool b a :
+    R (App (BVal b) a) (BVal b)
   | ProjAbs p a0 a1 :
     R a0 a1 ->
     R (Proj p (Abs a0)) (Abs (Proj p a1))
@@ -42,20 +42,19 @@ Module Par.
     R a0 a1 ->
     R b0 b1 ->
     R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
-  (* | ProjBool p b : *)
-  (*   R (Proj p (BVal b)) (BVal b) *)
-  (* | IfAbs a0 a1 b0 b1 c0 c1 d0 d1 : *)
-  (*   R a0 a1 -> *)
-  (*   R b0 b1 -> *)
-  (*   R c0 c1 -> *)
-  (*   R d0 d1 -> *)
-  (*   R (If (Abs a0) b0 c0) (Abs (If a0 (ren_Tm shift b0) (ren_Tm shift c0))) *)
-  (* | IfPair a0 a1 b0 b1 c0 c1 d0 d1 : *)
-  (*   R a0 a1 -> *)
-  (*   R b0 b1 -> *)
-  (*   R c0 c1 -> *)
-  (*   R d0 d1 -> *)
-  (*   R (If (Pair a0 b0) c0 d0) (Pair (If a0 c0 d0) (If b0 c0 d0)) *)
+  | ProjBool p b :
+    R (Proj p (BVal b)) (BVal b)
+  | IfAbs (a0 a1 : Tm (S n)) b0 b1 c0 c1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (Abs a0) b0 c0) (Abs (If a1 (ren_Tm shift b1) (ren_Tm shift c1)))
+  | IfPair a0 a1 b0 b1 c0 c1 d0 d1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R d0 d1 ->
+    R (If (Pair a0 b0) c0 d0) (Pair (If a1 c1 d1) (If b1 c1 d1))
   | IfBool a b0 b1 c0 c1 :
     R b0 b1 ->
     R c0 c1 ->
@@ -134,6 +133,14 @@ Module Par.
     R a0 b.
   Proof. move => ->; apply AppEta. Qed.
 
+  Lemma IfAbs' n (a0 a1  : Tm (S n)) (b0 b1 c0 c1 : Tm n)  u :
+    u = (Abs (If a1 (ren_Tm shift b1) (ren_Tm shift c1))) ->
+    @R (S n) a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (Abs a0) b0 c0) u.
+  Proof. move => ->. apply IfAbs. Qed.
+
   Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
     R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
   Proof.
@@ -142,11 +149,11 @@ Module Par.
     move => *; apply : AppAbs'; eauto; by asimpl.
     all : match goal with
           | [ |- context[var_zero]] => move => *; apply : AppEta'; eauto; by asimpl
+          | [ |- context[ren_Tm]] => move => * /=; apply : IfAbs'; eauto; by asimpl
           | _ => qauto ctrs:R use:ProjPair', IfBool'
           end.
   Qed.
 
-
   Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
     R a b -> R (subst_Tm ρ0 a) (subst_Tm ρ1 b).
@@ -157,8 +164,14 @@ Module Par.
       by asimpl.
       hauto l:on use:renaming inv:option.
     - hauto lq:on rew:off ctrs:R.
+    - hauto lq:on rew:off ctrs:R.
     - hauto l:on inv:option use:renaming ctrs:R.
     - hauto lq:on use:ProjPair'.
+    - hauto lq:on rew:off ctrs:R.
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ0 ρ1 hρ /=.
+      eapply IfAbs' with (a1 := (subst_Tm (up_Tm_Tm ρ1) a1)); eauto. by asimpl.
+      hauto l:on use:renaming inv:option.
+    - qauto l:on ctrs:R.
     - hauto q:on use:IfBool'.
     - move => n a0 a1 ha iha m ρ0 ρ1 hρ /=.
       apply : AppEta'; eauto. by asimpl.
@@ -204,6 +217,8 @@ Module Par.
       eexists. split.
       apply AppPair; hauto. subst.
       by asimpl.
+    - move => n b a m ξ []//=.
+      hauto q:on ctrs:R inv:Tm.
     - move => n p a0 a1 ha iha m ξ []//= p0 []//= t [*]. subst.
       spec_refl. move : iha => [b0 [? ?]]. subst.
       eexists. split. apply ProjAbs; eauto. by asimpl.
@@ -213,6 +228,23 @@ Module Par.
       move : ihb => [c0 [? ?]]. subst.
       eexists. split. by eauto using ProjPair.
       hauto q:on.
+    - move => n p b m ξ []//=.
+      move => + []//=.
+      hauto lq:on ctrs:R.
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ξ []//= t t0 t1 [h *]. subst.
+      case : t h => // t [*]. subst.
+      spec_refl.
+      move : iha => [a2 [iha ?]]. subst.
+      move : ihb => [b2 [ihb ?]]. subst.
+      move : ihc => [c2 [ihc ?]]. subst.
+      exists (Abs (If a2 (ren_Tm shift b2) (ren_Tm shift c2))).
+      split; last by asimpl.
+      hauto lq:on ctrs:R.
+    - move => n a0 a1 b0 b1 c0 c1 d0 d1 ha iha hb ihb hc ihc hd ihd m ξ []//=.
+      move => a2 b2 c2 [h *]. subst.
+      case : a2 h => //= a3 b3 [*]. subst.
+      spec_refl.
+      hauto lq:on ctrs:R.
     - move => n a b0 b1 c0 c1 hb ihb hc ihc m ξ []//= p0 p1 p2 [h *]. subst.
       spec_refl.
       move : ihb => [b0 [? ?]].
@@ -345,6 +377,8 @@ Module RPar.
     R b0 b1 ->
     R c0 c1 ->
     R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1))
+  | AppBool b a :
+    R (App (BVal b) a) (BVal b)
   | ProjAbs p a0 a1 :
     R a0 a1 ->
     R (Proj p (Abs a0)) (Abs (Proj p a1))
@@ -352,6 +386,19 @@ Module RPar.
     R a0 a1 ->
     R b0 b1 ->
     R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
+  | ProjBool p b :
+    R (Proj p (BVal b)) (BVal b)
+  | IfAbs (a0 a1 : Tm (S n)) b0 b1 c0 c1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (Abs a0) b0 c0) (Abs (If a1 (ren_Tm shift b1) (ren_Tm shift c1)))
+  | IfPair a0 a1 b0 b1 c0 c1 d0 d1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R d0 d1 ->
+    R (If (Pair a0 b0) c0 d0) (Pair (If a1 c1 d1) (If b1 c1 d1))
   | IfBool a b0 b1 c0 c1 :
     R b0 b1 ->
     R c0 c1 ->
@@ -421,13 +468,20 @@ Module RPar.
   Proof. move => ->. apply IfBool. Qed.
 
 
+  Lemma IfAbs' n (a0 a1  : Tm (S n)) (b0 b1 c0 c1 : Tm n)  u :
+    u = (Abs (If a1 (ren_Tm shift b1) (ren_Tm shift c1))) ->
+    @R (S n) a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If (Abs a0) b0 c0) u.
+  Proof. move => ->. apply IfAbs. Qed.
+
   Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
     R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
   Proof.
     move => h. move : m ξ.
-    elim : n a b /h.
-    move => *; apply : AppAbs'; eauto; by asimpl.
-    all : qauto ctrs:R use:ProjPair', IfBool'.
+    elim : n a b /h; try solve [(move => *; apply : AppAbs'; eauto; by asimpl) | (move => *; apply : IfAbs'; eauto; by asimpl)].
+    all : try qauto ctrs:R use:ProjPair', IfBool'.
   Qed.
 
   Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> Tm m) (ξ : fin m -> fin p) :
@@ -456,8 +510,14 @@ Module RPar.
       apply : AppAbs'; eauto using morphing_up.
       by asimpl.
     - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R use:ProjPair' use:morphing_up.
+    - sfirstorder.
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ0 ρ1 hρ /=.
+      apply : IfAbs'; eauto using morphing_up.
+      by asimpl.
+    - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R use:IfBool' use:morphing_up.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R use:morphing_up.
@@ -543,6 +603,9 @@ Module RPar.
       eexists. split.
       apply AppPair; hauto. subst.
       by asimpl.
+    - move => n b a m ρ hρ []//=; first by antiimp.
+      case => //=; first by antiimp.
+      hauto lq:on ctrs:R.
     - move => n p a0 a1 ha iha m ρ hρ []//=;
                first by antiimp.
       move => p0 []//= t [*]; first by antiimp. subst.
@@ -560,6 +623,31 @@ Module RPar.
       move : ihb => [c0 [? ?]]. subst.
       eexists. split. by eauto using ProjPair.
       hauto q:on.
+    - move => n p b m ρ hρ []//=; first by antiimp.
+      move => p0 + []//=. case => //=; first by antiimp.
+      hauto lq:on ctrs:R.
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ hρ []//=; first by antiimp.
+      move => + b2 c2 [+ *]. subst. case => //=; first by antiimp.
+      move => a [*]. subst.
+      have /var_or_bot_up hρ' := hρ.
+      move : iha hρ' => /[apply] iha.
+      move : ihb (hρ) => /[apply] ihb.
+      move : ihc hρ => /[apply] ihc.
+      spec_refl.
+      move : iha => [a0 [ha0 ?]]. subst.
+      move : ihb => [b0 [hb0 ?]]. subst.
+      move : ihc => [c0 [hc0 ?]]. subst.
+      exists (Abs (If a0 (ren_Tm shift b0) (ren_Tm shift c0))). split.
+      hauto lq:on ctrs:R.
+      by asimpl.
+    - move => n a0 a1 b0 b1 c0 c1 d0 d1 ha iha hb ihb hc ihc hd ihd m ρ hρ []//=; first by antiimp.
+      move => a2 b2 c2 [+ *]. subst. case : a2 => //=; first by antiimp.
+      move => a3 b3 [*]. subst.
+      have {}/iha := (hρ) => iha.
+      have {}/ihb := (hρ) => ihb.
+      have {}/ihc := (hρ) => ihc.
+      spec_refl.
+      hauto lq:on ctrs:R.
     - move => n a b0 b1 c0 c1 hb ihb hc ihc m ρ hρ []//=; first by antiimp.
       move => t t0 t1 [h *]. subst.
       case : t h => //=; first by antiimp.
@@ -631,332 +719,332 @@ Module RPar.
   Qed.
 End RPar.
 
-(***************** Beta rules only ***********************)
-Module RPar'.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
-  (***************** Beta ***********************)
-  | AppAbs a0 a1 b0 b1 :
-    R a0 a1 ->
-    R b0 b1 ->
-    R (App (Abs a0) b0)  (subst_Tm (scons b1 VarTm) a1)
-  | ProjPair p a0 a1 b0 b1 :
-    R a0 a1 ->
-    R b0 b1 ->
-    R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
+(* (***************** Beta rules only ***********************) *)
+(* Module RPar'. *)
+(*   Inductive R {n} : Tm n -> Tm n ->  Prop := *)
+(*   (***************** Beta ***********************) *)
+(*   | AppAbs a0 a1 b0 b1 : *)
+(*     R a0 a1 -> *)
+(*     R b0 b1 -> *)
+(*     R (App (Abs a0) b0)  (subst_Tm (scons b1 VarTm) a1) *)
+(*   | ProjPair p a0 a1 b0 b1 : *)
+(*     R a0 a1 -> *)
+(*     R b0 b1 -> *)
+(*     R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1) *)
 
 
-  (*************** Congruence ********************)
-  | Var i : R (VarTm i) (VarTm i)
-  | AbsCong a0 a1 :
-    R a0 a1 ->
-    R (Abs a0) (Abs a1)
-  | AppCong a0 a1 b0 b1 :
-    R a0 a1 ->
-    R b0 b1 ->
-    R (App a0 b0) (App a1 b1)
-  | PairCong a0 a1 b0 b1 :
-    R a0 a1 ->
-    R b0 b1 ->
-    R (Pair a0 b0) (Pair a1 b1)
-  | ProjCong p a0 a1 :
-    R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong p A0 A1 B0 B1:
-    R A0 A1 ->
-    R B0 B1 ->
-    R (TBind p A0 B0) (TBind p A1 B1)
-  | BotCong :
-    R Bot Bot
-  | UnivCong i :
-    R (Univ i) (Univ i).
+(*   (*************** Congruence ********************) *)
+(*   | Var i : R (VarTm i) (VarTm i) *)
+(*   | AbsCong a0 a1 : *)
+(*     R a0 a1 -> *)
+(*     R (Abs a0) (Abs a1) *)
+(*   | AppCong a0 a1 b0 b1 : *)
+(*     R a0 a1 -> *)
+(*     R b0 b1 -> *)
+(*     R (App a0 b0) (App a1 b1) *)
+(*   | PairCong a0 a1 b0 b1 : *)
+(*     R a0 a1 -> *)
+(*     R b0 b1 -> *)
+(*     R (Pair a0 b0) (Pair a1 b1) *)
+(*   | ProjCong p a0 a1 : *)
+(*     R a0 a1 -> *)
+(*     R (Proj p a0) (Proj p a1) *)
+(*   | BindCong p A0 A1 B0 B1: *)
+(*     R A0 A1 -> *)
+(*     R B0 B1 -> *)
+(*     R (TBind p A0 B0) (TBind p A1 B1) *)
+(*   | BotCong : *)
+(*     R Bot Bot *)
+(*   | UnivCong i : *)
+(*     R (Univ i) (Univ i). *)
 
-  Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
+(*   Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop. *)
 
-  Lemma refl n (a : Tm n) : R a a.
-  Proof.
-    induction a; hauto lq:on ctrs:R.
-  Qed.
+(*   Lemma refl n (a : Tm n) : R a a. *)
+(*   Proof. *)
+(*     induction a; hauto lq:on ctrs:R. *)
+(*   Qed. *)
 
-  Lemma AppAbs' n a0 a1 (b0 b1 t : Tm n) :
-    t = subst_Tm (scons b1 VarTm) a1 ->
-    R a0 a1 ->
-    R b0 b1 ->
-    R (App (Abs a0) b0) t.
-  Proof. move => ->. apply AppAbs. Qed.
+(*   Lemma AppAbs' n a0 a1 (b0 b1 t : Tm n) : *)
+(*     t = subst_Tm (scons b1 VarTm) a1 -> *)
+(*     R a0 a1 -> *)
+(*     R b0 b1 -> *)
+(*     R (App (Abs a0) b0) t. *)
+(*   Proof. move => ->. apply AppAbs. Qed. *)
 
-  Lemma ProjPair' n p (a0 a1 b0 b1 : Tm n) t :
-    t = (if p is PL then a1 else b1) ->
-    R a0 a1 ->
-    R b0 b1 ->
-    R (Proj p (Pair a0 b0)) t.
-  Proof.  move => > ->. apply ProjPair. Qed.
+(*   Lemma ProjPair' n p (a0 a1 b0 b1 : Tm n) t : *)
+(*     t = (if p is PL then a1 else b1) -> *)
+(*     R a0 a1 -> *)
+(*     R b0 b1 -> *)
+(*     R (Proj p (Pair a0 b0)) t. *)
+(*   Proof.  move => > ->. apply ProjPair. Qed. *)
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
-  Proof.
-    move => h. move : m ξ.
-    elim : n a b /h.
-    move => *; apply : AppAbs'; eauto; by asimpl.
-    all : qauto ctrs:R use:ProjPair'.
-  Qed.
+(*   Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) : *)
+(*     R a b -> R (ren_Tm ξ a) (ren_Tm ξ b). *)
+(*   Proof. *)
+(*     move => h. move : m ξ. *)
+(*     elim : n a b /h. *)
+(*     move => *; apply : AppAbs'; eauto; by asimpl. *)
+(*     all : qauto ctrs:R use:ProjPair'. *)
+(*   Qed. *)
 
-  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> Tm m) (ξ : fin m -> fin p) :
-    (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R ((funcomp (ren_Tm ξ) ρ0) i) ((funcomp (ren_Tm ξ) ρ1) i)).
-  Proof. eauto using renaming. Qed.
+(*   Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> Tm m) (ξ : fin m -> fin p) : *)
+(*     (forall i, R (ρ0 i) (ρ1 i)) -> *)
+(*     (forall i, R ((funcomp (ren_Tm ξ) ρ0) i) ((funcomp (ren_Tm ξ) ρ1) i)). *)
+(*   Proof. eauto using renaming. Qed. *)
 
-  Lemma morphing_ext n m (ρ0 ρ1 : fin n -> Tm m) a b  :
-    R a b ->
-    (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)).
-  Proof. hauto q:on inv:option. Qed.
+(*   Lemma morphing_ext n m (ρ0 ρ1 : fin n -> Tm m) a b  : *)
+(*     R a b -> *)
+(*     (forall i, R (ρ0 i) (ρ1 i)) -> *)
+(*     (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)). *)
+(*   Proof. hauto q:on inv:option. Qed. *)
 
-  Lemma morphing_up n m (ρ0 ρ1 : fin n -> Tm m) :
-    (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R (up_Tm_Tm ρ0 i) (up_Tm_Tm ρ1 i)).
-  Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_Tm_Tm. Qed.
+(*   Lemma morphing_up n m (ρ0 ρ1 : fin n -> Tm m) : *)
+(*     (forall i, R (ρ0 i) (ρ1 i)) -> *)
+(*     (forall i, R (up_Tm_Tm ρ0 i) (up_Tm_Tm ρ1 i)). *)
+(*   Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_Tm_Tm. Qed. *)
 
-  Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
-    (forall i, R (ρ0 i) (ρ1 i)) ->
-    R a b -> R (subst_Tm ρ0 a) (subst_Tm ρ1 b).
-  Proof.
-    move => + h. move : m ρ0 ρ1.
-    elim : n a b /h.
-    - move => *.
-      apply : AppAbs'; eauto using morphing_up.
-      by asimpl.
-    - hauto lq:on ctrs:R use:ProjPair' use:morphing_up.
-    - hauto lq:on ctrs:R use:morphing_up.
-    - hauto lq:on ctrs:R use:morphing_up.
-    - hauto lq:on ctrs:R use:morphing_up.
-    - hauto lq:on ctrs:R.
-    - hauto lq:on ctrs:R.
-    - hauto lq:on ctrs:R use:morphing_up.
-    - hauto lq:on ctrs:R.
-    - hauto lq:on ctrs:R.
-  Qed.
+(*   Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) : *)
+(*     (forall i, R (ρ0 i) (ρ1 i)) -> *)
+(*     R a b -> R (subst_Tm ρ0 a) (subst_Tm ρ1 b). *)
+(*   Proof. *)
+(*     move => + h. move : m ρ0 ρ1. *)
+(*     elim : n a b /h. *)
+(*     - move => *. *)
+(*       apply : AppAbs'; eauto using morphing_up. *)
+(*       by asimpl. *)
+(*     - hauto lq:on ctrs:R use:ProjPair' use:morphing_up. *)
+(*     - hauto lq:on ctrs:R use:morphing_up. *)
+(*     - hauto lq:on ctrs:R use:morphing_up. *)
+(*     - hauto lq:on ctrs:R use:morphing_up. *)
+(*     - hauto lq:on ctrs:R. *)
+(*     - hauto lq:on ctrs:R. *)
+(*     - hauto lq:on ctrs:R use:morphing_up. *)
+(*     - hauto lq:on ctrs:R. *)
+(*     - hauto lq:on ctrs:R. *)
+(*   Qed. *)
 
-  Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
-    R a b ->
-    R (subst_Tm ρ a) (subst_Tm ρ b).
-  Proof. hauto l:on use:morphing, refl. Qed.
+(*   Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) : *)
+(*     R a b -> *)
+(*     R (subst_Tm ρ a) (subst_Tm ρ b). *)
+(*   Proof. hauto l:on use:morphing, refl. Qed. *)
 
-  Lemma cong n (a b : Tm (S n)) c d :
-    R a b ->
-    R c d ->
-    R (subst_Tm (scons c VarTm) a) (subst_Tm (scons d VarTm) b).
-  Proof.
-    move => h0 h1. apply morphing => //=.
-    qauto l:on ctrs:R inv:option.
-  Qed.
+(*   Lemma cong n (a b : Tm (S n)) c d : *)
+(*     R a b -> *)
+(*     R c d -> *)
+(*     R (subst_Tm (scons c VarTm) a) (subst_Tm (scons d VarTm) b). *)
+(*   Proof. *)
+(*     move => h0 h1. apply morphing => //=. *)
+(*     qauto l:on ctrs:R inv:option. *)
+(*   Qed. *)
 
-  Lemma var_or_bot_imp {n} (a b : Tm n) :
-    var_or_bot a ->
-    a = b -> ~~ var_or_bot b -> False.
-  Proof.
-    hauto lq:on inv:Tm.
-  Qed.
+(*   Lemma var_or_bot_imp {n} (a b : Tm n) : *)
+(*     var_or_bot a -> *)
+(*     a = b -> ~~ var_or_bot b -> False. *)
+(*   Proof. *)
+(*     hauto lq:on inv:Tm. *)
+(*   Qed. *)
 
-  Lemma var_or_bot_up n m (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
-    (forall i, var_or_bot (up_Tm_Tm ρ i)).
-  Proof.
-    move => h /= [i|].
-    - asimpl.
-      move /(_ i) in h.
-      rewrite /funcomp.
-      move : (ρ i) h.
-      case => //=.
-    - sfirstorder.
-  Qed.
+(*   Lemma var_or_bot_up n m (ρ : fin n -> Tm m) : *)
+(*     (forall i, var_or_bot (ρ i)) -> *)
+(*     (forall i, var_or_bot (up_Tm_Tm ρ i)). *)
+(*   Proof. *)
+(*     move => h /= [i|]. *)
+(*     - asimpl. *)
+(*       move /(_ i) in h. *)
+(*       rewrite /funcomp. *)
+(*       move : (ρ i) h. *)
+(*       case => //=. *)
+(*     - sfirstorder. *)
+(*   Qed. *)
 
-  Local Ltac antiimp := qauto l:on use:var_or_bot_imp.
+(*   Local Ltac antiimp := qauto l:on use:var_or_bot_imp. *)
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
-    R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b.
-  Proof.
-    move E : (subst_Tm ρ a) => u hρ h.
-    move : n ρ hρ a E. elim : m u b/h.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
-               first by antiimp.
-      move => c c0 [+ ?]. subst.
-      case : c => //=; first by antiimp.
-      move => c [?]. subst.
-      spec_refl.
-      have /var_or_bot_up hρ' := hρ.
-      move : iha hρ' => /[apply] iha.
-      move : ihb hρ => /[apply] ihb.
-      spec_refl.
-      move : iha => [c1][ih0]?. subst.
-      move : ihb => [c2][ih1]?. subst.
-      eexists. split.
-      apply AppAbs; eauto.
-      by asimpl.
-    - move => n p a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
-               first by antiimp.
-      move => p0 []//=; first by antiimp. move => t t0[*].
-      subst.
-      have {}/iha := (hρ) => iha.
-      have {}/ihb := (hρ) => ihb.
-      spec_refl.
-      move : iha => [b0 [? ?]].
-      move : ihb => [c0 [? ?]]. subst.
-      eexists. split. by eauto using ProjPair.
-      hauto q:on.
-    - move => n i m ρ hρ []//=.
-      hauto l:on.
-    - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp.
-      move => t [*]. subst.
-      have /var_or_bot_up {}/iha := hρ => iha.
-      spec_refl.
-      move  :iha => [b0 [? ?]]. subst.
-      eexists. split. by apply AbsCong; eauto.
-      by asimpl.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
-               first by antiimp.
-      move => t t0 [*]. subst.
-      have {}/iha := (hρ) => iha.
-      have {}/ihb := (hρ) => ihb.
-      spec_refl.
-      move : iha => [b0 [? ?]]. subst.
-      move : ihb => [c0 [? ?]]. subst.
-      eexists. split. by apply AppCong; eauto.
-      done.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
-               first by antiimp.
-      move => t t0[*]. subst.
-      have {}/iha := (hρ) => iha.
-      have {}/ihb := (hρ) => ihb.
-      spec_refl.
-      move : iha => [b0 [? ?]]. subst.
-      move : ihb => [c0 [? ?]]. subst.
-      eexists. split. by apply PairCong; eauto.
-      by asimpl.
-    - move => n p a0 a1 ha iha m ρ hρ []//=;
-               first by antiimp.
-      move => p0 t [*]. subst.
-      have {}/iha := (hρ) => iha.
-      spec_refl.
-      move : iha => [b0 [? ?]]. subst.
-      eexists. split. apply ProjCong; eauto. reflexivity.
-    - move => n p A0 A1 B0 B1 ha iha hB ihB m ρ hρ []//=;
-               first by antiimp.
-      move => ? t t0 [*]. subst.
-      have {}/iha := (hρ) => iha.
-      have /var_or_bot_up {}/ihB := (hρ) => ihB.
-      spec_refl.
-      move : iha => [b0 [? ?]].
-      move : ihB => [c0 [? ?]]. subst.
-      eexists. split. by apply BindCong; eauto.
-      by asimpl.
-    - hauto q:on ctrs:R inv:Tm.
-    - move => n i n0 ρ hρ []//=; first by antiimp.
-      hauto l:on.
-  Qed.
-End RPar'.
+(*   Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) : *)
+(*     (forall i, var_or_bot (ρ i)) -> *)
+(*     R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b. *)
+(*   Proof. *)
+(*     move E : (subst_Tm ρ a) => u hρ h. *)
+(*     move : n ρ hρ a E. elim : m u b/h. *)
+(*     - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; *)
+(*                first by antiimp. *)
+(*       move => c c0 [+ ?]. subst. *)
+(*       case : c => //=; first by antiimp. *)
+(*       move => c [?]. subst. *)
+(*       spec_refl. *)
+(*       have /var_or_bot_up hρ' := hρ. *)
+(*       move : iha hρ' => /[apply] iha. *)
+(*       move : ihb hρ => /[apply] ihb. *)
+(*       spec_refl. *)
+(*       move : iha => [c1][ih0]?. subst. *)
+(*       move : ihb => [c2][ih1]?. subst. *)
+(*       eexists. split. *)
+(*       apply AppAbs; eauto. *)
+(*       by asimpl. *)
+(*     - move => n p a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; *)
+(*                first by antiimp. *)
+(*       move => p0 []//=; first by antiimp. move => t t0[*]. *)
+(*       subst. *)
+(*       have {}/iha := (hρ) => iha. *)
+(*       have {}/ihb := (hρ) => ihb. *)
+(*       spec_refl. *)
+(*       move : iha => [b0 [? ?]]. *)
+(*       move : ihb => [c0 [? ?]]. subst. *)
+(*       eexists. split. by eauto using ProjPair. *)
+(*       hauto q:on. *)
+(*     - move => n i m ρ hρ []//=. *)
+(*       hauto l:on. *)
+(*     - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp. *)
+(*       move => t [*]. subst. *)
+(*       have /var_or_bot_up {}/iha := hρ => iha. *)
+(*       spec_refl. *)
+(*       move  :iha => [b0 [? ?]]. subst. *)
+(*       eexists. split. by apply AbsCong; eauto. *)
+(*       by asimpl. *)
+(*     - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; *)
+(*                first by antiimp. *)
+(*       move => t t0 [*]. subst. *)
+(*       have {}/iha := (hρ) => iha. *)
+(*       have {}/ihb := (hρ) => ihb. *)
+(*       spec_refl. *)
+(*       move : iha => [b0 [? ?]]. subst. *)
+(*       move : ihb => [c0 [? ?]]. subst. *)
+(*       eexists. split. by apply AppCong; eauto. *)
+(*       done. *)
+(*     - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=; *)
+(*                first by antiimp. *)
+(*       move => t t0[*]. subst. *)
+(*       have {}/iha := (hρ) => iha. *)
+(*       have {}/ihb := (hρ) => ihb. *)
+(*       spec_refl. *)
+(*       move : iha => [b0 [? ?]]. subst. *)
+(*       move : ihb => [c0 [? ?]]. subst. *)
+(*       eexists. split. by apply PairCong; eauto. *)
+(*       by asimpl. *)
+(*     - move => n p a0 a1 ha iha m ρ hρ []//=; *)
+(*                first by antiimp. *)
+(*       move => p0 t [*]. subst. *)
+(*       have {}/iha := (hρ) => iha. *)
+(*       spec_refl. *)
+(*       move : iha => [b0 [? ?]]. subst. *)
+(*       eexists. split. apply ProjCong; eauto. reflexivity. *)
+(*     - move => n p A0 A1 B0 B1 ha iha hB ihB m ρ hρ []//=; *)
+(*                first by antiimp. *)
+(*       move => ? t t0 [*]. subst. *)
+(*       have {}/iha := (hρ) => iha. *)
+(*       have /var_or_bot_up {}/ihB := (hρ) => ihB. *)
+(*       spec_refl. *)
+(*       move : iha => [b0 [? ?]]. *)
+(*       move : ihB => [c0 [? ?]]. subst. *)
+(*       eexists. split. by apply BindCong; eauto. *)
+(*       by asimpl. *)
+(*     - hauto q:on ctrs:R inv:Tm. *)
+(*     - move => n i n0 ρ hρ []//=; first by antiimp. *)
+(*       hauto l:on. *)
+(*   Qed. *)
+(* End RPar'. *)
 
-Module ERed.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
-  (****************** Eta ***********************)
-  | AppEta a :
-    R a (Abs (App (ren_Tm shift a) (VarTm var_zero)))
-  | PairEta a :
-    R a (Pair (Proj PL a) (Proj PR a))
+(* Module ERed. *)
+(*   Inductive R {n} : Tm n -> Tm n ->  Prop := *)
+(*   (****************** Eta ***********************) *)
+(*   | AppEta a : *)
+(*     R a (Abs (App (ren_Tm shift a) (VarTm var_zero))) *)
+(*   | PairEta a : *)
+(*     R a (Pair (Proj PL a) (Proj PR a)) *)
 
-  (*************** Congruence ********************)
-  | AbsCong a0 a1 :
-    R a0 a1 ->
-    R (Abs a0) (Abs a1)
-  | AppCong0 a0 a1 b :
-    R a0 a1 ->
-    R (App a0 b) (App a1 b)
-  | AppCong1 a b0 b1 :
-    R b0 b1 ->
-    R (App a b0) (App a b1)
-  | PairCong0 a0 a1 b :
-    R a0 a1 ->
-    R (Pair a0 b) (Pair a1 b)
-  | PairCong1 a b0 b1 :
-    R b0 b1 ->
-    R (Pair a b0) (Pair a b1)
-  | ProjCong p a0 a1 :
-    R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong0 p A0 A1 B:
-    R A0 A1 ->
-    R (TBind p A0 B) (TBind p A1 B)
-  | BindCong1 p A B0 B1:
-    R B0 B1 ->
-    R (TBind p A B0) (TBind p A B1).
+(*   (*************** Congruence ********************) *)
+(*   | AbsCong a0 a1 : *)
+(*     R a0 a1 -> *)
+(*     R (Abs a0) (Abs a1) *)
+(*   | AppCong0 a0 a1 b : *)
+(*     R a0 a1 -> *)
+(*     R (App a0 b) (App a1 b) *)
+(*   | AppCong1 a b0 b1 : *)
+(*     R b0 b1 -> *)
+(*     R (App a b0) (App a b1) *)
+(*   | PairCong0 a0 a1 b : *)
+(*     R a0 a1 -> *)
+(*     R (Pair a0 b) (Pair a1 b) *)
+(*   | PairCong1 a b0 b1 : *)
+(*     R b0 b1 -> *)
+(*     R (Pair a b0) (Pair a b1) *)
+(*   | ProjCong p a0 a1 : *)
+(*     R a0 a1 -> *)
+(*     R (Proj p a0) (Proj p a1) *)
+(*   | BindCong0 p A0 A1 B: *)
+(*     R A0 A1 -> *)
+(*     R (TBind p A0 B) (TBind p A1 B) *)
+(*   | BindCong1 p A B0 B1: *)
+(*     R B0 B1 -> *)
+(*     R (TBind p A B0) (TBind p A B1). *)
 
-  Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
+(*   Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop. *)
 
-  Lemma AppEta' n a (u : Tm n) :
-    u = (Abs (App (ren_Tm shift a) (VarTm var_zero))) ->
-    R a u.
-  Proof. move => ->. apply AppEta. Qed.
+(*   Lemma AppEta' n a (u : Tm n) : *)
+(*     u = (Abs (App (ren_Tm shift a) (VarTm var_zero))) -> *)
+(*     R a u. *)
+(*   Proof. move => ->. apply AppEta. Qed. *)
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
-  Proof.
-    move => h. move : m ξ.
-    elim : n a b /h.
+(*   Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) : *)
+(*     R a b -> R (ren_Tm ξ a) (ren_Tm ξ b). *)
+(*   Proof. *)
+(*     move => h. move : m ξ. *)
+(*     elim : n a b /h. *)
 
-    move => n a m ξ.
-    apply AppEta'. by asimpl.
-    all : qauto ctrs:R.
-  Qed.
+(*     move => n a m ξ. *)
+(*     apply AppEta'. by asimpl. *)
+(*     all : qauto ctrs:R. *)
+(*   Qed. *)
 
-  Lemma substing n m (a : Tm n) b (ρ : fin n -> Tm m) :
-    R a b ->
-    R (subst_Tm ρ a) (subst_Tm ρ b).
-  Proof.
-    move => h. move : m ρ. elim : n a b / h => n.
-    move => a m ρ /=.
-    apply : AppEta'; eauto. by asimpl.
-    all : hauto ctrs:R inv:option use:renaming.
-  Qed.
+(*   Lemma substing n m (a : Tm n) b (ρ : fin n -> Tm m) : *)
+(*     R a b -> *)
+(*     R (subst_Tm ρ a) (subst_Tm ρ b). *)
+(*   Proof. *)
+(*     move => h. move : m ρ. elim : n a b / h => n. *)
+(*     move => a m ρ /=. *)
+(*     apply : AppEta'; eauto. by asimpl. *)
+(*     all : hauto ctrs:R inv:option use:renaming. *)
+(*   Qed. *)
 
-End ERed.
+(* End ERed. *)
 
-Module EReds.
+(* Module EReds. *)
 
-  #[local]Ltac solve_s_rec :=
-  move => *; eapply rtc_l; eauto;
-         hauto lq:on ctrs:ERed.R.
+(*   #[local]Ltac solve_s_rec := *)
+(*   move => *; eapply rtc_l; eauto; *)
+(*          hauto lq:on ctrs:ERed.R. *)
 
-  #[local]Ltac solve_s :=
-    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
+(*   #[local]Ltac solve_s := *)
+(*     repeat (induction 1; last by solve_s_rec); apply rtc_refl. *)
 
-  Lemma AbsCong n (a b : Tm (S n)) :
-    rtc ERed.R a b ->
-    rtc ERed.R (Abs a) (Abs b).
-  Proof. solve_s. Qed.
+(*   Lemma AbsCong n (a b : Tm (S n)) : *)
+(*     rtc ERed.R a b -> *)
+(*     rtc ERed.R (Abs a) (Abs b). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
-    rtc ERed.R a0 a1 ->
-    rtc ERed.R b0 b1 ->
-    rtc ERed.R (App a0 b0) (App a1 b1).
-  Proof. solve_s. Qed.
+(*   Lemma AppCong n (a0 a1 b0 b1 : Tm n) : *)
+(*     rtc ERed.R a0 a1 -> *)
+(*     rtc ERed.R b0 b1 -> *)
+(*     rtc ERed.R (App a0 b0) (App a1 b1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1 :
-    rtc ERed.R a0 a1 ->
-    rtc ERed.R b0 b1 ->
-    rtc ERed.R (TBind p a0 b0) (TBind p a1 b1).
-  Proof. solve_s. Qed.
+(*   Lemma BindCong n p (a0 a1 : Tm n) b0 b1 : *)
+(*     rtc ERed.R a0 a1 -> *)
+(*     rtc ERed.R b0 b1 -> *)
+(*     rtc ERed.R (TBind p a0 b0) (TBind p a1 b1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
-    rtc ERed.R a0 a1 ->
-    rtc ERed.R b0 b1 ->
-    rtc ERed.R (Pair a0 b0) (Pair a1 b1).
-  Proof. solve_s. Qed.
+(*   Lemma PairCong n (a0 a1 b0 b1 : Tm n) : *)
+(*     rtc ERed.R a0 a1 -> *)
+(*     rtc ERed.R b0 b1 -> *)
+(*     rtc ERed.R (Pair a0 b0) (Pair a1 b1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma ProjCong n p (a0 a1  : Tm n) :
-    rtc ERed.R a0 a1 ->
-    rtc ERed.R (Proj p a0) (Proj p a1).
-  Proof. solve_s. Qed.
-End EReds.
+(*   Lemma ProjCong n p (a0 a1  : Tm n) : *)
+(*     rtc ERed.R a0 a1 -> *)
+(*     rtc ERed.R (Proj p a0) (Proj p a1). *)
+(*   Proof. solve_s. Qed. *)
+(* End EReds. *)
 
 Module EPar.
   Inductive R {n} : Tm n -> Tm n ->  Prop :=
@@ -991,7 +1079,16 @@ Module EPar.
   | BotCong :
     R Bot Bot
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BoolCong :
+    R Bool Bool
+  | BValCong b :
+    R (BVal b) (BVal b)
+  | IfCong a0 a1 b0 b1 c0 c1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (If a0 b0 c0) (If a1 b1 c1).
 
   Lemma refl n (a : Tm n) : EPar.R a a.
   Proof.
@@ -1036,6 +1133,9 @@ Module EPar.
     - hauto l:on ctrs:R use:renaming inv:option.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
   Qed.
 
   Lemma substing n a0 a1 (b0 b1 : Tm n) :
@@ -1125,6 +1225,13 @@ Module RPars.
     rtc RPar.R (Pair a0 b0) (Pair a1 b1).
   Proof. solve_s. Qed.
 
+  Lemma IfCong n (a0 a1 b0 b1 c0 c1 : Tm n) :
+    rtc RPar.R a0 a1 ->
+    rtc RPar.R b0 b1 ->
+    rtc RPar.R c0 c1 ->
+    rtc RPar.R (If a0 b0 c0) (If a1 b1 c1).
+  Proof. solve_s. Qed.
+
   Lemma ProjCong n p (a0 a1  : Tm n) :
     rtc RPar.R a0 a1 ->
     rtc RPar.R (Proj p a0) (Proj p a1).
@@ -1180,92 +1287,92 @@ Module RPars.
 
 End RPars.
 
-Module RPars'.
+(* Module RPars'. *)
 
-  #[local]Ltac solve_s_rec :=
-  move => *; eapply rtc_l; eauto;
-         hauto lq:on ctrs:RPar'.R use:RPar'.refl.
+(*   #[local]Ltac solve_s_rec := *)
+(*   move => *; eapply rtc_l; eauto; *)
+(*          hauto lq:on ctrs:RPar'.R use:RPar'.refl. *)
 
-  #[local]Ltac solve_s :=
-    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
+(*   #[local]Ltac solve_s := *)
+(*     repeat (induction 1; last by solve_s_rec); apply rtc_refl. *)
 
-  Lemma AbsCong n (a b : Tm (S n)) :
-    rtc RPar'.R a b ->
-    rtc RPar'.R (Abs a) (Abs b).
-  Proof. solve_s. Qed.
+(*   Lemma AbsCong n (a b : Tm (S n)) : *)
+(*     rtc RPar'.R a b -> *)
+(*     rtc RPar'.R (Abs a) (Abs b). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R b0 b1 ->
-    rtc RPar'.R (App a0 b0) (App a1 b1).
-  Proof. solve_s. Qed.
+(*   Lemma AppCong n (a0 a1 b0 b1 : Tm n) : *)
+(*     rtc RPar'.R a0 a1 -> *)
+(*     rtc RPar'.R b0 b1 -> *)
+(*     rtc RPar'.R (App a0 b0) (App a1 b1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1 :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R b0 b1 ->
-    rtc RPar'.R (TBind p a0 b0) (TBind p a1 b1).
-  Proof. solve_s. Qed.
+(*   Lemma BindCong n p (a0 a1 : Tm n) b0 b1 : *)
+(*     rtc RPar'.R a0 a1 -> *)
+(*     rtc RPar'.R b0 b1 -> *)
+(*     rtc RPar'.R (TBind p a0 b0) (TBind p a1 b1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R b0 b1 ->
-    rtc RPar'.R (Pair a0 b0) (Pair a1 b1).
-  Proof. solve_s. Qed.
+(*   Lemma PairCong n (a0 a1 b0 b1 : Tm n) : *)
+(*     rtc RPar'.R a0 a1 -> *)
+(*     rtc RPar'.R b0 b1 -> *)
+(*     rtc RPar'.R (Pair a0 b0) (Pair a1 b1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma ProjCong n p (a0 a1  : Tm n) :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R (Proj p a0) (Proj p a1).
-  Proof. solve_s. Qed.
+(*   Lemma ProjCong n p (a0 a1  : Tm n) : *)
+(*     rtc RPar'.R a0 a1 -> *)
+(*     rtc RPar'.R (Proj p a0) (Proj p a1). *)
+(*   Proof. solve_s. Qed. *)
 
-  Lemma renaming n (a0 a1 : Tm n) m (ξ : fin n -> fin m) :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R (ren_Tm ξ a0) (ren_Tm ξ a1).
-  Proof.
-    induction 1.
-    - apply rtc_refl.
-    - eauto using RPar'.renaming, rtc_l.
-  Qed.
+(*   Lemma renaming n (a0 a1 : Tm n) m (ξ : fin n -> fin m) : *)
+(*     rtc RPar'.R a0 a1 -> *)
+(*     rtc RPar'.R (ren_Tm ξ a0) (ren_Tm ξ a1). *)
+(*   Proof. *)
+(*     induction 1. *)
+(*     - apply rtc_refl. *)
+(*     - eauto using RPar'.renaming, rtc_l. *)
+(*   Qed. *)
 
-  Lemma weakening n (a0 a1 : Tm n) :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R (ren_Tm shift a0) (ren_Tm shift a1).
-  Proof. apply renaming. Qed.
+(*   Lemma weakening n (a0 a1 : Tm n) : *)
+(*     rtc RPar'.R a0 a1 -> *)
+(*     rtc RPar'.R (ren_Tm shift a0) (ren_Tm shift a1). *)
+(*   Proof. apply renaming. Qed. *)
 
-  Lemma Abs_inv n (a : Tm (S n)) b :
-    rtc RPar'.R (Abs a) b -> exists a', b = Abs a' /\ rtc RPar'.R a a'.
-  Proof.
-    move E : (Abs a) => b0 h. move : a E.
-    elim : b0 b / h.
-    - hauto lq:on ctrs:rtc.
-    - hauto lq:on ctrs:rtc inv:RPar'.R, rtc.
-  Qed.
+(*   Lemma Abs_inv n (a : Tm (S n)) b : *)
+(*     rtc RPar'.R (Abs a) b -> exists a', b = Abs a' /\ rtc RPar'.R a a'. *)
+(*   Proof. *)
+(*     move E : (Abs a) => b0 h. move : a E. *)
+(*     elim : b0 b / h. *)
+(*     - hauto lq:on ctrs:rtc. *)
+(*     - hauto lq:on ctrs:rtc inv:RPar'.R, rtc. *)
+(*   Qed. *)
 
-  Lemma morphing n m (a b : Tm n) (ρ : fin n -> Tm m) :
-    rtc RPar'.R a b ->
-    rtc RPar'.R (subst_Tm ρ a) (subst_Tm ρ b).
-  Proof. induction 1; qauto l:on ctrs:rtc use:RPar'.substing. Qed.
+(*   Lemma morphing n m (a b : Tm n) (ρ : fin n -> Tm m) : *)
+(*     rtc RPar'.R a b -> *)
+(*     rtc RPar'.R (subst_Tm ρ a) (subst_Tm ρ b). *)
+(*   Proof. induction 1; qauto l:on ctrs:rtc use:RPar'.substing. Qed. *)
 
-  Lemma substing n (a b : Tm (S n)) c :
-    rtc RPar'.R a b ->
-    rtc RPar'.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b).
-  Proof. hauto lq:on use:morphing inv:option. Qed.
+(*   Lemma substing n (a b : Tm (S n)) c : *)
+(*     rtc RPar'.R a b -> *)
+(*     rtc RPar'.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b). *)
+(*   Proof. hauto lq:on use:morphing inv:option. Qed. *)
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
-    rtc RPar'.R (subst_Tm ρ a) b -> exists b0, rtc RPar'.R a b0 /\ subst_Tm ρ b0 = b.
-  Proof.
-    move E  :(subst_Tm ρ a) => u hρ h.
-    move : a E.
-    elim : u b /h.
-    - sfirstorder.
-    - move => a b c h0 h1 ih1 a0 ?. subst.
-      move /RPar'.antirenaming : h0.
-      move /(_ hρ).
-      move => [b0 [h2 ?]]. subst.
-      hauto lq:on rew:off ctrs:rtc.
-  Qed.
+(*   Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) : *)
+(*     (forall i, var_or_bot (ρ i)) -> *)
+(*     rtc RPar'.R (subst_Tm ρ a) b -> exists b0, rtc RPar'.R a b0 /\ subst_Tm ρ b0 = b. *)
+(*   Proof. *)
+(*     move E  :(subst_Tm ρ a) => u hρ h. *)
+(*     move : a E. *)
+(*     elim : u b /h. *)
+(*     - sfirstorder. *)
+(*     - move => a b c h0 h1 ih1 a0 ?. subst. *)
+(*       move /RPar'.antirenaming : h0. *)
+(*       move /(_ hρ). *)
+(*       move => [b0 [h2 ?]]. subst. *)
+(*       hauto lq:on rew:off ctrs:rtc. *)
+(*   Qed. *)
 
-End RPars'.
+(* End RPars'. *)
 
 Lemma Abs_EPar n a (b : Tm n) :
   EPar.R (Abs a) b ->
@@ -1421,6 +1528,7 @@ Proof.
       split.
       hauto lq:on use:relations.rtc_transitive, RPars.AppCong.
       apply EPar.PairCong; by apply EPar.AppCong.
+    + admit.
     + hauto lq:on ctrs:EPar.R use:RPars.AppCong.
   - hauto lq:on ctrs:EPar.R inv:RPar.R use:RPars.PairCong.
   - move => n p a b0 h0 ih0 b1.
@@ -1435,11 +1543,21 @@ Proof.
       move /Pair_EPar : ih1 => [/(_ p)[d [ihd ihd']] _].
       exists d. split => //.
       hauto lq:on use:RPars.ProjCong, relations.rtc_transitive.
+    + admit.
     + hauto lq:on ctrs:EPar.R use:RPars.ProjCong.
   - hauto lq:on inv:RPar.R ctrs:EPar.R, rtc use:RPars.BindCong.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
-Qed.
+  - hauto l:on ctrs:EPar.R inv:RPar.R.
+  - hauto l:on ctrs:EPar.R inv:RPar.R.
+  - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc u.
+    elim /RPar.inv => //= _.
+    + admit.
+    + admit.
+    + admit.
+    + (* Need the rule that if commutes with abs *)
+      admit.
+Admitted.
 
 Lemma commutativity1 n (a b0 b1 : Tm n) :
   EPar.R a b0 -> rtc RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ EPar.R b1 c.
@@ -1615,7 +1733,10 @@ Proof.
     exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
   - qauto use:Bot_EPar', EPar.refl.
   - qauto use:Univ_EPar', EPar.refl.
-Qed.
+  - admit.
+  - admit.
+  - admit.
+Admitted.
 
 Function tstar {n} (a : Tm n) :=
   match a with

From d0760cd9dbb99910918cebd24a29b9abe6ddb985 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sat, 11 Jan 2025 00:07:30 -0500
Subject: [PATCH 04/13] Stuck at commutativity (trouble with if pair)

---
 theories/fp_red.v | 48 ++++++++++++++++++++++++++++++++++++++++++++---
 1 file changed, 45 insertions(+), 3 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index a8c8a3a..5094c57 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1374,6 +1374,11 @@ End RPars.
 
 (* End RPars'. *)
 
+(* (forall p, exists d, rtc RPar.R (Proj p c) d /\ EPar.R (if p is PL then a else b) d) *)
+(* (exists d0 d1, rtc RPar.R (App (ren_Tm shift c) (VarTm var_zero)) *)
+(*                 (Pair (App (ren_Tm shift d0) (VarTm var_zero))(App (ren_Tm shift d1) (VarTm var_zero))) /\ *)
+(*     EPar.R a d0 /\ EPar.R b d1) *)
+
 Lemma Abs_EPar n a (b : Tm n) :
   EPar.R (Abs a) b ->
   (exists d, EPar.R a d /\
@@ -1426,6 +1431,32 @@ Proof.
       apply rtc_once. apply : RPar.ProjAbs; eauto using RPar.refl.
 Qed.
 
+Lemma Abs_EPar_If n (a : Tm (S n)) q :
+  EPar.R (Abs a) q ->
+  exists d, EPar.R a d /\
+         forall b c, rtc RPar.R (If q b c) (Abs (If d (ren_Tm shift b) (ren_Tm shift c))).
+Proof.
+  move E : (Abs a) => u h.
+  move : a E.
+  elim : n u q /h => //= n.
+  - move => a0 a1 ha _ b ?. subst.
+    move /Abs_EPar : ha.
+    move => [[d [h0 h1]] _].
+    exists d.
+    split => // b0 c.
+    apply : rtc_l.
+    apply RPar.IfAbs; auto using RPar.refl.
+    apply RPars.AbsCong.
+    apply RPars.IfCong; auto using rtc_refl.
+  - move => a0 a1 ha _ a ?. subst.
+    move /Abs_EPar : ha => [_ [d [h0 h1]]].
+    exists d. split => //.
+    move => b c.
+    apply : rtc_l.
+    apply RPar.IfPair; auto using RPar.refl.
+Admitted.
+
+
 Lemma Pair_EPar n (a b c : Tm n) :
   EPar.R (Pair a b) c ->
   (forall p, exists d, rtc RPar.R (Proj p c) d /\ EPar.R (if p is PL then a else b) d) /\
@@ -1528,7 +1559,9 @@ Proof.
       split.
       hauto lq:on use:relations.rtc_transitive, RPars.AppCong.
       apply EPar.PairCong; by apply EPar.AppCong.
-    + admit.
+    + case : a0 ha iha => //=.
+      move => b ha hc b3 a0 [*]. subst.
+      admit.
     + hauto lq:on ctrs:EPar.R use:RPars.AppCong.
   - hauto lq:on ctrs:EPar.R inv:RPar.R use:RPars.PairCong.
   - move => n p a b0 h0 ih0 b1.
@@ -1543,7 +1576,8 @@ Proof.
       move /Pair_EPar : ih1 => [/(_ p)[d [ihd ihd']] _].
       exists d. split => //.
       hauto lq:on use:RPars.ProjCong, relations.rtc_transitive.
-    + admit.
+    + move => p0 b2 [*]. subst.
+      admit.
     + hauto lq:on ctrs:EPar.R use:RPars.ProjCong.
   - hauto lq:on inv:RPar.R ctrs:EPar.R, rtc use:RPars.BindCong.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
@@ -1552,7 +1586,15 @@ Proof.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
   - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc u.
     elim /RPar.inv => //= _.
-    + admit.
+    + move => a2 a3 b2 b3 c2 c3 ha2 hb2 hc2 [*]. subst.
+      move /(_ _ ltac:(by eauto using RPar.AbsCong)) : iha => [c [ih0 ih1]].
+      move /Abs_EPar : ih1 => [[d [ih2 ih3]]  _].
+      have {}/ihb := hb2. move => [b4 [ihb0 ihb1]].
+      have {}/ihc := hc2. move => [c4 [ihc0 ihc1]].
+
+      exists (Abs (If d (ren_Tm shift b4) (ren_Tm shift c4))).
+      split. admit.
+      hauto lq:on ctrs:EPar.R use:EPar.renaming.
     + admit.
     + admit.
     + (* Need the rule that if commutes with abs *)

From f3718707f2b4668ee61ec48531a6576dbf29b123 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 20 Jan 2025 13:56:40 -0500
Subject: [PATCH 05/13] Add constants to the reduction semantics

---
 syntax.sig                   |   6 +-
 theories/Autosubst2/syntax.v |  29 +++---
 theories/fp_red.v            | 169 ++++++++++++-----------------------
 3 files changed, 74 insertions(+), 130 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index 4549f7f..d268994 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -3,14 +3,14 @@ Tm(VarTm) : Type
 PTag : Type
 TTag : Type
 
-TPi : TTag
-TSig : TTag
 PL : PTag
 PR : PTag
+TPi : TTag
+TSig : TTag
 Abs : (bind Tm in Tm) -> Tm
 App : Tm -> Tm -> Tm
 Pair : Tm -> Tm -> Tm
 Proj : PTag -> Tm -> Tm
 TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm
-Bot : Tm
+Const : TTag -> Tm
 Univ : nat -> Tm
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index b972476..05fb668 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -40,7 +40,7 @@ Inductive Tm (n_Tm : nat) : Type :=
   | Pair : Tm n_Tm -> Tm n_Tm -> Tm n_Tm
   | Proj : PTag -> Tm n_Tm -> Tm n_Tm
   | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
-  | Bot : Tm n_Tm
+  | Const : TTag -> Tm n_Tm
   | Univ : nat -> Tm n_Tm.
 
 Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)}
@@ -83,9 +83,10 @@ exact (eq_trans
          (ap (fun x => TBind m_Tm t0 t1 x) H2)).
 Qed.
 
-Lemma congr_Bot {m_Tm : nat} : Bot m_Tm = Bot m_Tm.
+Lemma congr_Const {m_Tm : nat} {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
+  Const m_Tm s0 = Const m_Tm t0.
 Proof.
-exact (eq_refl).
+exact (eq_trans eq_refl (ap (fun x => Const m_Tm x) H0)).
 Qed.
 
 Lemma congr_Univ {m_Tm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
@@ -116,7 +117,7 @@ Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
   | Proj _ s0 s1 => Proj n_Tm s0 (ren_Tm xi_Tm s1)
   | TBind _ s0 s1 s2 =>
       TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
-  | Bot _ => Bot n_Tm
+  | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
   end.
 
@@ -143,7 +144,7 @@ Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
   | Proj _ s0 s1 => Proj n_Tm s0 (subst_Tm sigma_Tm s1)
   | TBind _ s0 s1 s2 =>
       TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
-  | Bot _ => Bot n_Tm
+  | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
   end.
 
@@ -183,7 +184,7 @@ subst_Tm sigma_Tm s = s :=
   | TBind _ s0 s1 s2 =>
       congr_TBind (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
         (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -227,7 +228,7 @@ Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
       congr_TBind (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
         (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upExtRen_Tm_Tm _ _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -272,7 +273,7 @@ Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
       congr_TBind (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
         (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm)
            s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -317,7 +318,7 @@ Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
       congr_TBind (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
         (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -373,7 +374,7 @@ Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -450,7 +451,7 @@ ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
         (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -528,7 +529,7 @@ subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -644,7 +645,7 @@ Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
       congr_TBind (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
         (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm)
            (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2)
-  | Bot _ => congr_Bot
+  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   end.
 
@@ -846,7 +847,7 @@ Arguments VarTm {n_Tm}.
 
 Arguments Univ {n_Tm}.
 
-Arguments Bot {n_Tm}.
+Arguments Const {n_Tm}.
 
 Arguments TBind {n_Tm}.
 
diff --git a/theories/fp_red.v b/theories/fp_red.v
index d2fbdf6..30f3354 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -69,9 +69,8 @@ Module Par.
     R A0 A1 ->
     R B0 B1 ->
     R (TBind p A0 B0) (TBind p A1 B1)
-  (* Bot is useful for making the prov function computable  *)
-  | BotCong :
-    R Bot Bot
+  | ConstCong k :
+    R (Const k) (Const k)
   | UnivCong i :
     R (Univ i) (Univ i).
 
@@ -212,7 +211,7 @@ Module Par.
       move : ihB => [c0 [? ?]]. subst.
       eexists. split. by apply BindCong; eauto.
       by asimpl.
-    - move => n n0 ξ []//=. hauto l:on.
+    - move => n k m ξ []//=. hauto l:on.
     - move => n i n0 ξ []//=. hauto l:on.
   Qed.
 
@@ -275,10 +274,10 @@ Module Pars.
 
 End Pars.
 
-Definition var_or_bot {n} (a : Tm n) :=
+Definition var_or_const {n} (a : Tm n) :=
   match a with
   | VarTm _ => true
-  | Bot => true
+  | Const _ => true
   | _ => false
   end.
 
@@ -324,8 +323,8 @@ Module RPar.
     R A0 A1 ->
     R B0 B1 ->
     R (TBind p A0 B0) (TBind p A1 B1)
-  | BotCong :
-    R Bot Bot
+  | ConstCong k :
+    R (Const k) (Const k)
   | UnivCong i :
     R (Univ i) (Univ i).
 
@@ -571,8 +570,8 @@ Module RPar'.
     R A0 A1 ->
     R B0 B1 ->
     R (TBind p A0 B0) (TBind p A1 B1)
-  | BotCong :
-    R Bot Bot
+  | ConstCong k :
+    R (Const k) (Const k)
   | UnivCong i :
     R (Univ i) (Univ i).
 
@@ -656,16 +655,16 @@ Module RPar'.
     qauto l:on ctrs:R inv:option.
   Qed.
 
-  Lemma var_or_bot_imp {n} (a b : Tm n) :
-    var_or_bot a ->
-    a = b -> ~~ var_or_bot b -> False.
+  Lemma var_or_const_imp {n} (a b : Tm n) :
+    var_or_const a ->
+    a = b -> ~~ var_or_const b -> False.
   Proof.
     hauto lq:on inv:Tm.
   Qed.
 
-  Lemma var_or_bot_up n m (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
-    (forall i, var_or_bot (up_Tm_Tm ρ i)).
+  Lemma var_or_const_up n m (ρ : fin n -> Tm m) :
+    (forall i, var_or_const (ρ i)) ->
+    (forall i, var_or_const (up_Tm_Tm ρ i)).
   Proof.
     move => h /= [i|].
     - asimpl.
@@ -676,10 +675,10 @@ Module RPar'.
     - sfirstorder.
   Qed.
 
-  Local Ltac antiimp := qauto l:on use:var_or_bot_imp.
+  Local Ltac antiimp := qauto l:on use:var_or_const_imp.
 
   Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
+    (forall i, var_or_const (ρ i)) ->
     R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b.
   Proof.
     move E : (subst_Tm ρ a) => u hρ h.
@@ -690,7 +689,7 @@ Module RPar'.
       case : c => //=; first by antiimp.
       move => c [?]. subst.
       spec_refl.
-      have /var_or_bot_up hρ' := hρ.
+      have /var_or_const_up hρ' := hρ.
       move : iha hρ' => /[apply] iha.
       move : ihb hρ => /[apply] ihb.
       spec_refl.
@@ -714,7 +713,7 @@ Module RPar'.
       hauto l:on.
     - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp.
       move => t [*]. subst.
-      have /var_or_bot_up {}/iha := hρ => iha.
+      have /var_or_const_up {}/iha := hρ => iha.
       spec_refl.
       move  :iha => [b0 [? ?]]. subst.
       eexists. split. by apply AbsCong; eauto.
@@ -750,7 +749,7 @@ Module RPar'.
                first by antiimp.
       move => ? t t0 [*]. subst.
       have {}/iha := (hρ) => iha.
-      have /var_or_bot_up {}/ihB := (hρ) => ihB.
+      have /var_or_const_up {}/ihB := (hρ) => ihB.
       spec_refl.
       move : iha => [b0 [? ?]].
       move : ihB => [c0 [? ?]]. subst.
@@ -894,8 +893,8 @@ Module EPar.
     R A0 A1 ->
     R B0 B1 ->
     R (TBind p A0 B0) (TBind p A1 B1)
-  | BotCong :
-    R Bot Bot
+  | ConstCong k :
+    R (Const k) (Const k)
   | UnivCong i :
     R (Univ i) (Univ i).
 
@@ -1070,7 +1069,7 @@ Module RPars.
   Proof. hauto lq:on use:morphing inv:option. Qed.
 
   Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
+    (forall i, var_or_const (ρ i)) ->
     rtc RPar.R (subst_Tm ρ a) b -> exists b0, rtc RPar.R a b0 /\ subst_Tm ρ b0 = b.
   Proof.
     move E  :(subst_Tm ρ a) => u hρ h.
@@ -1157,7 +1156,7 @@ Module RPars'.
   Proof. hauto lq:on use:morphing inv:option. Qed.
 
   Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
+    (forall i, var_or_const (ρ i)) ->
     rtc RPar'.R (subst_Tm ρ a) b -> exists b0, rtc RPar'.R a b0 /\ subst_Tm ρ b0 = b.
   Proof.
     move E  :(subst_Tm ρ a) => u hρ h.
@@ -1444,15 +1443,15 @@ Proof.
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma Bot_EPar' n (u : Tm n) :
-  EPar.R Bot u ->
-  rtc OExp.R Bot u.
-  move E : Bot => t h.
-  move : E. elim : n t u /h => //=.
-  - move => n a0 a1 h ih ?. subst.
+Lemma Const_EPar' n k (u : Tm n) :
+  EPar.R (Const k) u ->
+  rtc OExp.R (Const k) u.
+  move E : (Const k) => t h.
+  move : k E. elim : n t u /h => //=.
+  - move => n a0 a1 h ih k ?. subst.
     specialize ih with (1 := eq_refl).
     hauto lq:on ctrs:OExp.R use:rtc_r.
-  - move => n a0 a1 h ih ?. subst.
+  - move => n a0 a1 h ih k ?. subst.
     specialize ih with (1 := eq_refl).
     hauto lq:on ctrs:OExp.R use:rtc_r.
   - hauto l:on ctrs:OExp.R.
@@ -1519,7 +1518,7 @@ Proof.
     move : OExp.commutativity0 h2; repeat move/[apply].
     move => [d h].
     exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
-  - qauto use:Bot_EPar', EPar.refl.
+  - qauto use:Const_EPar', EPar.refl.
   - qauto use:Univ_EPar', EPar.refl.
 Qed.
 
@@ -1536,7 +1535,7 @@ Function tstar {n} (a : Tm n) :=
   | Proj p (Abs a) => (Abs (Proj p (tstar a)))
   | Proj p a => Proj p (tstar a)
   | TBind p a b => TBind p (tstar a) (tstar b)
-  | Bot => Bot
+  | Const k => Const k
   | Univ i => Univ i
   end.
 
@@ -1568,7 +1567,7 @@ Function tstar' {n} (a : Tm n) :=
   | Proj p (Pair a b) => if p is PL then (tstar' a) else (tstar' b)
   | Proj p a => Proj p (tstar' a)
   | TBind p a b => TBind p (tstar' a) (tstar' b)
-  | Bot => Bot
+  | Const k => Const k
   | Univ i => Univ i
   end.
 
@@ -1616,63 +1615,6 @@ Proof.
   sfirstorder use:relations.diamond_confluent, EPar_diamond.
 Qed.
 
-Fixpoint depth_tm {n} (a : Tm n) :=
-  match a with
-  | VarTm _ => 1
-  | TBind _ A B => 1 + max (depth_tm A) (depth_tm B)
-  | Abs a => 1 + depth_tm a
-  | App a b => 1 + max (depth_tm a) (depth_tm b)
-  | Proj p a => 1 + depth_tm a
-  | Pair a b => 1 + max (depth_tm a) (depth_tm b)
-  | Bot => 1
-  | Univ i => 1
-  end.
-
-Lemma depth_ren n m (ξ: fin n -> fin m) a :
-  depth_tm a = depth_tm (ren_Tm ξ a).
-Proof.
-  move : m ξ. elim : n / a; scongruence.
-Qed.
-
-Lemma depth_subst n m (ρ : fin n -> Tm m) a :
-  (forall i, depth_tm (ρ i) = 1) ->
-  depth_tm a = depth_tm (subst_Tm ρ a).
-Proof.
-  move : m ρ. elim : n / a.
-  - sfirstorder.
-  - move => n a iha m ρ hρ.
-    simpl.
-    f_equal. apply iha.
-    destruct i as [i|].
-    + simpl.
-      by rewrite -depth_ren.
-    + by simpl.
-  - hauto lq:on rew:off.
-  - hauto lq:on rew:off.
-  - hauto lq:on rew:off.
-  - move => n p a iha b ihb m ρ hρ.
-    simpl. f_equal.
-    f_equal.
-    by apply iha.
-    apply ihb.
-    destruct i as [i|].
-    + simpl.
-      by rewrite -depth_ren.
-    + by simpl.
-  - sfirstorder.
-  - sfirstorder.
-Qed.
-
-Lemma depth_subst_bool n (a : Tm (S n)) :
-  depth_tm a = depth_tm (subst_Tm (scons Bot VarTm) a).
-Proof.
-  apply depth_subst.
-  destruct i as [i|] => //=.
-Qed.
-
-Local Ltac prov_tac := sfirstorder use:depth_ren.
-Local Ltac extract_tac := rewrite -?depth_subst_bool;hauto use:depth_subst_bool.
-
 Definition prov_bind {n} p0 A0 B0 (a : Tm n)  :=
   match a with
   | TBind p A B => p = p0 /\ rtc Par.R A A0 /\ rtc Par.R B B0
@@ -1704,8 +1646,8 @@ Inductive prov {n} : Tm n -> Tm n -> Prop :=
 | P_Proj h p a :
   prov h a ->
   prov h (Proj p a)
-| P_Bot  :
-  prov Bot Bot
+| P_Const k  :
+  prov (Const k) (Const k)
 | P_Var i :
   prov (VarTm i) (VarTm i)
 | P_Univ i :
@@ -1789,7 +1731,7 @@ Proof.
   split.
   move => h. constructor. move => b. asimpl. by constructor.
   inversion 1; subst.
-  specialize H2 with (b := Bot).
+  specialize H2 with (b := Const TPi).
   move : H2. asimpl. inversion 1; subst. done.
 Qed.
 
@@ -1845,11 +1787,11 @@ Qed.
 Fixpoint extract {n} (a : Tm n) : Tm n :=
   match a with
   | TBind p A B => TBind p A B
-  | Abs a => subst_Tm (scons Bot VarTm) (extract a)
+  | Abs a => subst_Tm (scons (Const TPi) VarTm) (extract a)
   | App a b => extract a
   | Pair a b => extract a
   | Proj p a => extract a
-  | Bot => Bot
+  | Const k => Const k
   | VarTm i => VarTm i
   | Univ i => Univ i
   end.
@@ -1886,7 +1828,7 @@ Proof.
 Qed.
 
 Lemma ren_subst_bot n (a : Tm (S n)) :
-  extract (subst_Tm (scons Bot VarTm) a) = subst_Tm (scons Bot VarTm) (extract a).
+  extract (subst_Tm (scons (Const TPi) VarTm) a) = subst_Tm (scons (Const TPi) VarTm) (extract a).
 Proof.
   apply ren_morphing. destruct i as [i|] => //=.
 Qed.
@@ -1896,7 +1838,7 @@ Definition prov_extract_spec {n} u (a : Tm n) :=
   | TBind p A B => exists A0 B0, extract a = TBind p A0 B0 /\ rtc Par.R A A0 /\ rtc Par.R B B0
   | Univ i => extract a = Univ i
   | VarTm i => extract a = VarTm i
-  | Bot => extract a = Bot
+  | (Const TPi) => extract a = (Const TPi)
   | _ => True
   end.
 
@@ -1909,22 +1851,22 @@ Proof.
   - move => h a ha ih.
     case : h ha ih => //=.
     + move => i ha ih.
-      move /(_ Bot) in ih.
+      move /(_ (Const TPi)) in ih.
       rewrite -ih.
       by rewrite ren_subst_bot.
     + move => p A B h ih.
-      move /(_ Bot) : ih => [A0][B0][h0][h1]h2.
+      move /(_ (Const TPi)) : ih => [A0][B0][h0][h1]h2.
       rewrite ren_subst_bot in h0.
       rewrite h0.
       eauto.
-    + move => _ /(_ Bot).
+    + move => p _ /(_ (Const TPi)).
       by rewrite ren_subst_bot.
-    + move => i h /(_ Bot).
+    + move => i h /(_ (Const TPi)).
       by rewrite ren_subst_bot => ->.
   - hauto lq:on.
   - hauto lq:on.
   - hauto lq:on.
-  - sfirstorder.
+  - case => //=.
   - sfirstorder.
   - sfirstorder.
 Qed.
@@ -2400,23 +2342,24 @@ Fixpoint ne {n} (a : Tm n) :=
   match a with
   | VarTm i => true
   | TBind _ A B => false
-  | Bot => true
   | App a b => ne a && nf b
   | Abs a => false
   | Univ _ => false
   | Proj _ a => ne a
   | Pair _ _ => false
+  | Const _ => false
   end
 with nf {n} (a : Tm n) :=
   match a with
   | VarTm i => true
   | TBind _ A B => nf A && nf B
-  | Bot => true
+  | (Const TPi) => true
   | App a b => ne a && nf b
   | Abs a => nf a
   | Univ _ => true
   | Proj _ a => ne a
   | Pair a b => nf a && nf b
+  | Const _ => true
 end.
 
 Lemma ne_nf n a : @ne n a -> nf a.
@@ -2482,15 +2425,15 @@ Create HintDb nfne.
 #[export]Hint Resolve nf_wn ne_nf wne_wn nf_refl : nfne.
 
 Lemma ne_nf_antiren n m (a : Tm n) (ρ : fin n -> Tm m) :
-  (forall i, var_or_bot (ρ i)) ->
+  (forall i, var_or_const (ρ i)) ->
   (ne (subst_Tm ρ a) -> ne a) /\ (nf (subst_Tm ρ a) -> nf a).
 Proof.
   move : m ρ. elim : n / a => //;
-   hauto b:on drew:off use:RPar.var_or_bot_up.
+   hauto b:on drew:off use:RPar.var_or_const_up.
 Qed.
 
 Lemma wn_antirenaming n m a (ρ : fin n -> Tm m) :
-  (forall i, var_or_bot (ρ i)) ->
+  (forall i, var_or_const (ρ i)) ->
   wn (subst_Tm ρ a) -> wn a.
 Proof.
   rewrite /wn => hρ.
@@ -2503,10 +2446,10 @@ Proof.
 Qed.
 
 Lemma ext_wn n (a : Tm n) :
-    wn (App a Bot) ->
+    wn (App a (Const TPi)) ->
     wn a.
 Proof.
-  move E : (App a Bot) => a0 [v [hr hv]].
+  move E : (App a (Const TPi)) => a0 [v [hr hv]].
   move : a E.
   move : hv.
   elim : a0 v / hr.
@@ -2517,9 +2460,9 @@ Proof.
     case : a0 hr0=>// => b0 b1.
     elim /RPar'.inv=>// _.
     + move => a0 a3 b2 b3 ? ? [? ?] ? [? ?]. subst.
-      have ? : b3 = Bot by hauto lq:on inv:RPar'.R. subst.
+      have ? : b3 = (Const TPi) by hauto lq:on inv:RPar'.R. subst.
       suff : wn (Abs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
-      have : wn (subst_Tm (scons Bot VarTm) a3) by sfirstorder.
+      have : wn (subst_Tm (scons (Const TPi) VarTm) a3) by sfirstorder.
       move => h. apply wn_abs.
       move : h. apply wn_antirenaming.
       hauto lq:on rew:off inv:option.

From d68df5d0bc84617900727ed91a1948c360ec066b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 20 Jan 2025 15:33:46 -0500
Subject: [PATCH 06/13] Add compile

---
 theories/compile.v | 136 +++++++++++++++++++++++++++++++++++++++++++++
 theories/fp_red.v  |  49 +++++++++++-----
 2 files changed, 172 insertions(+), 13 deletions(-)
 create mode 100644 theories/compile.v

diff --git a/theories/compile.v b/theories/compile.v
new file mode 100644
index 0000000..d9d6ca3
--- /dev/null
+++ b/theories/compile.v
@@ -0,0 +1,136 @@
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red.
+Require Import ssreflect ssrbool.
+From Hammer Require Import Tactics.
+
+Module Compile.
+  Fixpoint F {n} (a : Tm n) : Tm n :=
+    match a with
+    | TBind p A B => Pair (Pair (Const p) (F A)) (Abs (F B))
+    | Const k => Const k
+    | Univ i => Univ i
+    | Abs a => Abs (F a)
+    | App a b => App (F a) (F b)
+    | VarTm i => VarTm i
+    | Pair a b => Pair (F a) (F b)
+    | Proj t a => Proj t (F a)
+    end.
+
+  Lemma renaming n m (a : Tm n) (ξ : fin n -> fin m) :
+    F (ren_Tm ξ a)= ren_Tm ξ (F a).
+  Proof. move : m ξ. elim : n  / a => //=; scongruence. Qed.
+
+  #[local]Hint Rewrite Compile.renaming : compile.
+
+  Lemma morphing n m (a : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
+    (forall i, ρ0 i = F (ρ1 i)) ->
+    subst_Tm ρ0 (F a) = F (subst_Tm ρ1 a).
+  Proof.
+    move : m ρ0 ρ1. elim : n / a => n//=.
+    - hauto lq:on inv:option rew:db:compile unfold:funcomp.
+    - hauto lq:on rew:off.
+    - hauto lq:on rew:off.
+    - hauto lq:on.
+    - hauto lq:on inv:option rew:db:compile unfold:funcomp.
+  Qed.
+
+  Lemma substing n b (a : Tm (S n)) :
+    subst_Tm (scons (F b) VarTm) (F a) = F (subst_Tm (scons b VarTm) a).
+  Proof.
+    apply morphing.
+    case => //=.
+  Qed.
+
+End Compile.
+
+#[export] Hint Rewrite Compile.renaming Compile.morphing : compile.
+
+
+Module Join.
+  Definition R {n} (a b : Tm n) := join (Compile.F a) (Compile.F b).
+
+  Lemma BindInj n p0 p1 (A0 A1 : Tm n) B0 B1  :
+    R (TBind p0 A0 B0) (TBind p1 A1 B1) -> p0 = p1 /\ R A0 A1 /\ R B0 B1.
+  Proof.
+    rewrite /R /= !join_pair_inj.
+    move => [[/join_const_inj h0 h1] h2].
+    apply abs_eq in h2.
+    evar (t : Tm (S n)).
+    have : join (App (ren_Tm shift (Abs (Compile.F B1))) (VarTm var_zero)) t by
+      apply Join.FromPar; apply Par.AppAbs; auto using Par.refl.
+    subst t. rewrite -/ren_Tm.
+    move : h2. move /join_transitive => /[apply]. asimpl => h2.
+    tauto.
+  Qed.
+
+  Lemma UnivInj n i j : R (Univ i : Tm n) (Univ j) -> i = j.
+  Proof. hauto l:on use:join_univ_inj. Qed.
+
+End Join.
+
+Module Equiv.
+  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  (***************** Beta ***********************)
+  | AppAbs a b :
+    R (App (Abs a) b)  (subst_Tm (scons b VarTm) a)
+  | ProjPair p a b :
+    R (Proj p (Pair a b)) (if p is PL then a else b)
+
+  (****************** Eta ***********************)
+  | AppEta a :
+    R a (Abs (App (ren_Tm shift a) (VarTm var_zero)))
+  | PairEta a :
+    R a (Pair (Proj PL a) (Proj PR a))
+
+  (*************** Congruence ********************)
+  | Var i : R (VarTm i) (VarTm i)
+  | AbsCong a b :
+    R a b ->
+    R (Abs a) (Abs b)
+  | AppCong a0 a1 b0 b1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R (App a0 b0) (App a1 b1)
+  | PairCong a0 a1 b0 b1 :
+    R a0 a1 ->
+    R b0 b1 ->
+    R (Pair a0 b0) (Pair a1 b1)
+  | ProjCong p a0 a1 :
+    R a0 a1 ->
+    R (Proj p a0) (Proj p a1)
+  | BindCong p A0 A1 B0 B1:
+    R A0 A1 ->
+    R B0 B1 ->
+    R (TBind p A0 B0) (TBind p A1 B1)
+  | UnivCong i :
+    R (Univ i) (Univ i).
+End Equiv.
+
+Module EquivJoin.
+  Lemma FromEquiv n (a b : Tm n) : Equiv.R a b -> Join.R a b.
+  Proof.
+    move => h. elim : n a b /h => n.
+    - move => a b.
+      rewrite /Join.R /join /=.
+      eexists. split. apply relations.rtc_once.
+      apply Par.AppAbs; auto using Par.refl.
+      rewrite Compile.substing.
+      apply relations.rtc_refl.
+    - move => p a b.
+      apply Join.FromPar.
+      simpl. apply : Par.ProjPair'; auto using Par.refl.
+      case : p => //=.
+    - move => a. apply Join.FromPar => /=.
+      apply : Par.AppEta'; auto using Par.refl.
+      by autorewrite with compile.
+    - move => a. apply Join.FromPar => /=.
+      apply : Par.PairEta; auto using Par.refl.
+    - hauto l:on use:Join.FromPar, Par.Var.
+    - hauto lq:on use:Join.AbsCong.
+    - qauto l:on use:Join.AppCong.
+    - qauto l:on use:Join.PairCong.
+    - qauto use:Join.ProjCong.
+    - rewrite /Join.R => p A0 A1 B0 B1 _ hA _ hB /=.
+      sfirstorder use:Join.PairCong,Join.AbsCong,Join.FromPar,Par.ConstCong.
+    - hauto l:on.
+  Qed.
+End EquivJoin.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 30f3354..adb05ea 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -410,16 +410,16 @@ Module RPar.
     qauto l:on ctrs:R inv:option.
   Qed.
 
-  Lemma var_or_bot_imp {n} (a b : Tm n) :
-    var_or_bot a ->
-    a = b -> ~~ var_or_bot b -> False.
+  Lemma var_or_const_imp {n} (a b : Tm n) :
+    var_or_const a ->
+    a = b -> ~~ var_or_const b -> False.
   Proof.
     hauto lq:on inv:Tm.
   Qed.
 
-  Lemma var_or_bot_up n m (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
-    (forall i, var_or_bot (up_Tm_Tm ρ i)).
+  Lemma var_or_const_up n m (ρ : fin n -> Tm m) :
+    (forall i, var_or_const (ρ i)) ->
+    (forall i, var_or_const (up_Tm_Tm ρ i)).
   Proof.
     move => h /= [i|].
     - asimpl.
@@ -430,10 +430,10 @@ Module RPar.
     - sfirstorder.
   Qed.
 
-  Local Ltac antiimp := qauto l:on use:var_or_bot_imp.
+  Local Ltac antiimp := qauto l:on use:var_or_const_imp.
 
   Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
-    (forall i, var_or_bot (ρ i)) ->
+    (forall i, var_or_const (ρ i)) ->
     R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b.
   Proof.
     move E : (subst_Tm ρ a) => u hρ h.
@@ -444,7 +444,7 @@ Module RPar.
       case : c => //=; first by antiimp.
       move => c [?]. subst.
       spec_refl.
-      have /var_or_bot_up hρ' := hρ.
+      have /var_or_const_up hρ' := hρ.
       move : iha hρ' => /[apply] iha.
       move : ihb hρ => /[apply] ihb.
       spec_refl.
@@ -470,7 +470,7 @@ Module RPar.
     - move => n p a0 a1 ha iha m ρ hρ []//=;
                first by antiimp.
       move => p0 []//= t [*]; first by antiimp. subst.
-      have  /var_or_bot_up {}/iha  := hρ => iha.
+      have  /var_or_const_up {}/iha  := hρ => iha.
       spec_refl. move : iha => [b0 [? ?]]. subst.
       eexists. split. apply ProjAbs; eauto. by asimpl.
     - move => n p a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
@@ -488,7 +488,7 @@ Module RPar.
       hauto l:on.
     - move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp.
       move => t [*]. subst.
-      have /var_or_bot_up {}/iha := hρ => iha.
+      have /var_or_const_up {}/iha := hρ => iha.
       spec_refl.
       move  :iha => [b0 [? ?]]. subst.
       eexists. split. by apply AbsCong; eauto.
@@ -524,7 +524,7 @@ Module RPar.
                first by antiimp.
       move => ? t t0 [*]. subst.
       have {}/iha := (hρ) => iha.
-      have /var_or_bot_up {}/ihB := (hρ) => ihB.
+      have /var_or_const_up {}/ihB := (hρ) => ihB.
       spec_refl.
       move : iha => [b0 [? ?]].
       move : ihB => [c0 [? ?]]. subst.
@@ -1838,7 +1838,7 @@ Definition prov_extract_spec {n} u (a : Tm n) :=
   | TBind p A B => exists A0 B0, extract a = TBind p A0 B0 /\ rtc Par.R A A0 /\ rtc Par.R B B0
   | Univ i => extract a = Univ i
   | VarTm i => extract a = VarTm i
-  | (Const TPi) => extract a = (Const TPi)
+  | (Const i) => extract a = (Const i)
   | _ => True
   end.
 
@@ -2250,6 +2250,15 @@ Proof.
   apply prov_extract.
 Qed.
 
+Lemma pars_const_inv n i (c : Tm n) :
+  rtc Par.R (Const i) c ->
+  extract c = Const i.
+Proof.
+  have : prov (Const i) (Const i : Tm n) by sfirstorder.
+  move : prov_pars. repeat move/[apply].
+  apply prov_extract.
+Qed.
+
 Lemma pars_pi_inv n p (A : Tm n) B C :
   rtc Par.R (TBind p A B) C ->
   exists A0 B0, extract C = TBind p A0 B0 /\
@@ -2277,6 +2286,14 @@ Proof.
   sauto l:on use:pars_univ_inv.
 Qed.
 
+Lemma pars_const_inj n i j (C : Tm n) :
+  rtc Par.R (Const i) C ->
+  rtc Par.R (Const j) C ->
+  i = j.
+Proof.
+  sauto l:on use:pars_const_inv.
+Qed.
+
 Lemma pars_pi_inj n p0 p1 (A0 A1 : Tm n) B0 B1 C :
   rtc Par.R (TBind p0 A0 B0) C ->
   rtc Par.R (TBind p1 A1 B1) C ->
@@ -2314,6 +2331,12 @@ Proof.
   sfirstorder use:pars_univ_inj.
 Qed.
 
+Lemma join_const_inj n i j :
+  join (Const i : Tm n) (Const j) -> i = j.
+Proof.
+  sfirstorder use:pars_const_inj.
+Qed.
+
 Lemma join_pi_inj n p0 p1 (A0 A1 : Tm n) B0 B1 :
   join (TBind p0 A0 B0) (TBind p1 A1 B1) ->
   p0 = p1 /\ join A0 A1 /\ join B0 B1.

From 9c9ce52b63ed216299646e519e38dfb419b31ed7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 20 Jan 2025 16:28:44 -0500
Subject: [PATCH 07/13] Recover the contra lemmas

---
 syntax.sig                   |  1 +
 theories/Autosubst2/syntax.v | 20 ++++++++-
 theories/compile.v           |  1 +
 theories/fp_red.v            | 83 ++++++++++++++++++++++++++++--------
 theories/logrel.v            | 33 +++++++++-----
 5 files changed, 110 insertions(+), 28 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index d268994..3101cab 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -14,3 +14,4 @@ Proj : PTag -> Tm -> Tm
 TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm
 Const : TTag -> Tm
 Univ : nat -> Tm
+Bot : Tm
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index 05fb668..a5cb002 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -41,7 +41,8 @@ Inductive Tm (n_Tm : nat) : Type :=
   | Proj : PTag -> Tm n_Tm -> Tm n_Tm
   | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
   | Const : TTag -> Tm n_Tm
-  | Univ : nat -> Tm n_Tm.
+  | Univ : nat -> Tm n_Tm
+  | Bot : Tm n_Tm.
 
 Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)}
   (H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0.
@@ -95,6 +96,11 @@ Proof.
 exact (eq_trans eq_refl (ap (fun x => Univ m_Tm x) H0)).
 Qed.
 
+Lemma congr_Bot {m_Tm : nat} : Bot m_Tm = Bot m_Tm.
+Proof.
+exact (eq_refl).
+Qed.
+
 Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) :
   fin (S m) -> fin (S n).
 Proof.
@@ -119,6 +125,7 @@ Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
       TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
   | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
+  | Bot _ => Bot n_Tm
   end.
 
 Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) :
@@ -146,6 +153,7 @@ Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
       TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
   | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
+  | Bot _ => Bot n_Tm
   end.
 
 Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm)
@@ -186,6 +194,7 @@ subst_Tm sigma_Tm s = s :=
         (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n)
@@ -230,6 +239,7 @@ Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
            (upExtRen_Tm_Tm _ _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm)
@@ -275,6 +285,7 @@ Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
            s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
@@ -320,6 +331,7 @@ Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
            (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat}
@@ -376,6 +388,7 @@ Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
            (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@@ -453,6 +466,7 @@ ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
            (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@@ -531,6 +545,7 @@ subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
            (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
@@ -647,6 +662,7 @@ Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
            (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
+  | Bot _ => congr_Bot
   end.
 
 Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
@@ -845,6 +861,8 @@ Core.
 
 Arguments VarTm {n_Tm}.
 
+Arguments Bot {n_Tm}.
+
 Arguments Univ {n_Tm}.
 
 Arguments Const {n_Tm}.
diff --git a/theories/compile.v b/theories/compile.v
index d9d6ca3..3eaaf42 100644
--- a/theories/compile.v
+++ b/theories/compile.v
@@ -13,6 +13,7 @@ Module Compile.
     | VarTm i => VarTm i
     | Pair a b => Pair (F a) (F b)
     | Proj t a => Proj t (F a)
+    | Bot => Bot
     end.
 
   Lemma renaming n m (a : Tm n) (ξ : fin n -> fin m) :
diff --git a/theories/fp_red.v b/theories/fp_red.v
index adb05ea..6932f84 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -72,7 +72,9 @@ Module Par.
   | ConstCong k :
     R (Const k) (Const k)
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BotCong :
+    R Bot Bot.
 
   Lemma refl n (a : Tm n) : R a a.
     elim : n /a; hauto ctrs:R.
@@ -134,6 +136,7 @@ Module Par.
     - hauto l:on inv:option ctrs:R use:renaming.
     - sfirstorder.
     - sfirstorder.
+    - sfirstorder.
   Qed.
 
   Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
@@ -213,6 +216,7 @@ Module Par.
       by asimpl.
     - move => n k m ξ []//=. hauto l:on.
     - move => n i n0 ξ []//=. hauto l:on.
+    - hauto q:on inv:Tm ctrs:R.
   Qed.
 
 End Par.
@@ -277,7 +281,7 @@ End Pars.
 Definition var_or_const {n} (a : Tm n) :=
   match a with
   | VarTm _ => true
-  | Const _ => true
+  | Bot => true
   | _ => false
   end.
 
@@ -326,7 +330,9 @@ Module RPar.
   | ConstCong k :
     R (Const k) (Const k)
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BotCong :
+    R Bot Bot.
 
   Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
 
@@ -394,6 +400,7 @@ Module RPar.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
   Qed.
 
   Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
@@ -533,6 +540,7 @@ Module RPar.
     - hauto q:on ctrs:R inv:Tm.
     - move => n i n0 ρ hρ []//=; first by antiimp.
       hauto l:on.
+    - move => n m ρ hρ []//=; hauto lq:on ctrs:R.
   Qed.
 End RPar.
 
@@ -573,7 +581,9 @@ Module RPar'.
   | ConstCong k :
     R (Const k) (Const k)
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BotCong :
+    R Bot Bot.
 
   Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
 
@@ -639,6 +649,7 @@ Module RPar'.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
   Qed.
 
   Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
@@ -758,6 +769,7 @@ Module RPar'.
     - hauto q:on ctrs:R inv:Tm.
     - move => n i n0 ρ hρ []//=; first by antiimp.
       hauto l:on.
+    - hauto q:on inv:Tm ctrs:R.
   Qed.
 End RPar'.
 
@@ -896,7 +908,9 @@ Module EPar.
   | ConstCong k :
     R (Const k) (Const k)
   | UnivCong i :
-    R (Univ i) (Univ i).
+    R (Univ i) (Univ i)
+  | BotCong :
+    R Bot Bot.
 
   Lemma refl n (a : Tm n) : EPar.R a a.
   Proof.
@@ -941,6 +955,7 @@ Module EPar.
     - hauto l:on ctrs:R use:renaming inv:option.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
+    - hauto lq:on ctrs:R.
   Qed.
 
   Lemma substing n a0 a1 (b0 b1 : Tm n) :
@@ -1344,6 +1359,7 @@ Proof.
   - hauto lq:on inv:RPar.R ctrs:EPar.R, rtc use:RPars.BindCong.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
+  - hauto l:on ctrs:EPar.R inv:RPar.R.
 Qed.
 
 Lemma commutativity1 n (a b0 b1 : Tm n) :
@@ -1457,6 +1473,20 @@ Lemma Const_EPar' n k (u : Tm n) :
   - hauto l:on ctrs:OExp.R.
 Qed.
 
+Lemma Bot_EPar' n (u : Tm n) :
+  EPar.R (Bot) u ->
+  rtc OExp.R (Bot) u.
+  move E : (Bot) => t h.
+  move : E. elim : n t u /h => //=.
+  - move => n a0 a1 h ih ?. subst.
+    specialize ih with (1 := eq_refl).
+    hauto lq:on ctrs:OExp.R use:rtc_r.
+  - move => n a0 a1 h ih ?. subst.
+    specialize ih with (1 := eq_refl).
+    hauto lq:on ctrs:OExp.R use:rtc_r.
+  - hauto l:on ctrs:OExp.R.
+Qed.
+
 Lemma Univ_EPar' n i (u : Tm n) :
   EPar.R (Univ i) u ->
   rtc OExp.R (Univ i) u.
@@ -1520,6 +1550,7 @@ Proof.
     exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
   - qauto use:Const_EPar', EPar.refl.
   - qauto use:Univ_EPar', EPar.refl.
+  - qauto use:Bot_EPar', EPar.refl.
 Qed.
 
 Function tstar {n} (a : Tm n) :=
@@ -1537,6 +1568,7 @@ Function tstar {n} (a : Tm n) :=
   | TBind p a b => TBind p (tstar a) (tstar b)
   | Const k => Const k
   | Univ i => Univ i
+  | Bot => Bot
   end.
 
 Lemma RPar_triangle n (a : Tm n) : forall b, RPar.R a b -> RPar.R b (tstar a).
@@ -1555,6 +1587,7 @@ Proof.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
+  - hauto lq:on inv:RPar.R ctrs:RPar.R.
 Qed.
 
 Function tstar' {n} (a : Tm n) :=
@@ -1569,6 +1602,7 @@ Function tstar' {n} (a : Tm n) :=
   | TBind p a b => TBind p (tstar' a) (tstar' b)
   | Const k => Const k
   | Univ i => Univ i
+  | Bot => Bot
   end.
 
 Lemma RPar'_triangle n (a : Tm n) : forall b, RPar'.R a b -> RPar'.R b (tstar' a).
@@ -1585,6 +1619,7 @@ Proof.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
+  - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
 Qed.
 
 Lemma RPar_diamond n (c a1 b1 : Tm n) :
@@ -1651,7 +1686,9 @@ Inductive prov {n} : Tm n -> Tm n -> Prop :=
 | P_Var i :
   prov (VarTm i) (VarTm i)
 | P_Univ i :
-  prov (Univ i) (Univ i).
+  prov (Univ i) (Univ i)
+| P_Bot :
+  prov Bot Bot.
 
 Lemma ERed_EPar n (a b : Tm n) : ERed.R a b -> EPar.R a b.
 Proof.
@@ -1671,6 +1708,7 @@ Proof.
   - eauto using EReds.BindCong.
   - auto using rtc_refl.
   - auto using rtc_refl.
+  - auto using rtc_refl.
 Qed.
 
 Lemma EPar_Par n (a b : Tm n) : EPar.R a b -> Par.R a b.
@@ -1723,6 +1761,7 @@ Proof.
   - hauto lq:on ctrs:prov inv:RPar.R.
   - hauto l:on ctrs:RPar.R inv:RPar.R.
   - hauto l:on ctrs:RPar.R inv:RPar.R.
+  - hauto l:on ctrs:RPar.R inv:RPar.R.
 Qed.
 
 
@@ -1777,6 +1816,7 @@ Proof.
   - hauto lq:on inv:ERed.R, prov ctrs:prov.
   - hauto lq:on inv:ERed.R, prov ctrs:prov.
   - hauto lq:on inv:ERed.R, prov ctrs:prov.
+  - hauto lq:on inv:ERed.R, prov ctrs:prov.
 Qed.
 
 Lemma prov_ereds n (u : Tm n) a b : prov u a -> rtc ERed.R a b -> prov u b.
@@ -1787,13 +1827,14 @@ Qed.
 Fixpoint extract {n} (a : Tm n) : Tm n :=
   match a with
   | TBind p A B => TBind p A B
-  | Abs a => subst_Tm (scons (Const TPi) VarTm) (extract a)
+  | Abs a => subst_Tm (scons Bot VarTm) (extract a)
   | App a b => extract a
   | Pair a b => extract a
   | Proj p a => extract a
   | Const k => Const k
   | VarTm i => VarTm i
   | Univ i => Univ i
+  | Bot => Bot
   end.
 
 Lemma ren_extract n m (a : Tm n) (ξ : fin n -> fin m) :
@@ -1810,6 +1851,7 @@ Proof.
   - hauto q:on.
   - sfirstorder.
   - sfirstorder.
+  - sfirstorder.
 Qed.
 
 Lemma ren_morphing n m (a : Tm n) (ρ : fin n -> Tm m) :
@@ -1828,7 +1870,7 @@ Proof.
 Qed.
 
 Lemma ren_subst_bot n (a : Tm (S n)) :
-  extract (subst_Tm (scons (Const TPi) VarTm) a) = subst_Tm (scons (Const TPi) VarTm) (extract a).
+  extract (subst_Tm (scons Bot VarTm) a) = subst_Tm (scons Bot VarTm) (extract a).
 Proof.
   apply ren_morphing. destruct i as [i|] => //=.
 Qed.
@@ -1839,6 +1881,7 @@ Definition prov_extract_spec {n} u (a : Tm n) :=
   | Univ i => extract a = Univ i
   | VarTm i => extract a = VarTm i
   | (Const i) => extract a = (Const i)
+  | Bot => extract a = Bot
   | _ => True
   end.
 
@@ -1851,24 +1894,28 @@ Proof.
   - move => h a ha ih.
     case : h ha ih => //=.
     + move => i ha ih.
-      move /(_ (Const TPi)) in ih.
+      move /(_ Bot) in ih.
       rewrite -ih.
       by rewrite ren_subst_bot.
     + move => p A B h ih.
-      move /(_ (Const TPi)) : ih => [A0][B0][h0][h1]h2.
+      move /(_ Bot) : ih => [A0][B0][h0][h1]h2.
       rewrite ren_subst_bot in h0.
       rewrite h0.
       eauto.
-    + move => p _ /(_ (Const TPi)).
+    + move => p _ /(_ Bot).
       by rewrite ren_subst_bot.
-    + move => i h /(_ (Const TPi)).
+    + move => i h /(_ Bot).
       by rewrite ren_subst_bot => ->.
+    + move /(_ Bot).
+      move => h /(_ Bot).
+      by rewrite ren_subst_bot.
   - hauto lq:on.
   - hauto lq:on.
   - hauto lq:on.
   - case => //=.
   - sfirstorder.
   - sfirstorder.
+  - sfirstorder.
 Qed.
 
 Definition union {A : Type} (R0 R1 : A -> A -> Prop) a b :=
@@ -2072,6 +2119,7 @@ Proof.
   - sfirstorder use:ERPars.BindCong.
   - sfirstorder.
   - sfirstorder.
+  - sfirstorder.
 Qed.
 
 Lemma Pars_ERPar n (a b : Tm n) : rtc Par.R a b -> rtc ERPar.R a b.
@@ -2371,18 +2419,19 @@ Fixpoint ne {n} (a : Tm n) :=
   | Proj _ a => ne a
   | Pair _ _ => false
   | Const _ => false
+  | Bot => true
   end
 with nf {n} (a : Tm n) :=
   match a with
   | VarTm i => true
   | TBind _ A B => nf A && nf B
-  | (Const TPi) => true
   | App a b => ne a && nf b
   | Abs a => nf a
   | Univ _ => true
   | Proj _ a => ne a
   | Pair a b => nf a && nf b
   | Const _ => true
+  | Bot => true
 end.
 
 Lemma ne_nf n a : @ne n a -> nf a.
@@ -2469,10 +2518,10 @@ Proof.
 Qed.
 
 Lemma ext_wn n (a : Tm n) :
-    wn (App a (Const TPi)) ->
+    wn (App a Bot) ->
     wn a.
 Proof.
-  move E : (App a (Const TPi)) => a0 [v [hr hv]].
+  move E : (App a (Bot)) => a0 [v [hr hv]].
   move : a E.
   move : hv.
   elim : a0 v / hr.
@@ -2483,9 +2532,9 @@ Proof.
     case : a0 hr0=>// => b0 b1.
     elim /RPar'.inv=>// _.
     + move => a0 a3 b2 b3 ? ? [? ?] ? [? ?]. subst.
-      have ? : b3 = (Const TPi) by hauto lq:on inv:RPar'.R. subst.
+      have ? : b3 = (Bot) by hauto lq:on inv:RPar'.R. subst.
       suff : wn (Abs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
-      have : wn (subst_Tm (scons (Const TPi) VarTm) a3) by sfirstorder.
+      have : wn (subst_Tm (scons (Bot) VarTm) a3) by sfirstorder.
       move => h. apply wn_abs.
       move : h. apply wn_antirenaming.
       hauto lq:on rew:off inv:option.
diff --git a/theories/logrel.v b/theories/logrel.v
index a4b15d2..a421cda 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1,5 +1,4 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
-Require Import fp_red.
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red compile.
 From Hammer Require Import Tactics.
 From Equations Require Import Equations.
 Require Import ssreflect ssrbool.
@@ -295,9 +294,9 @@ Proof.
   - hauto lq:on ctrs:prov.
   - hauto lq:on rew:off ctrs:prov b:on.
   - hauto lq:on ctrs:prov.
-  - move => n.
+  - move => n k.
     have : @prov n Bot Bot by auto using P_Bot.
-    tauto.
+    eauto.
 Qed.
 
 Lemma ne_pars_inv n (a b : Tm n) :
@@ -311,23 +310,37 @@ Lemma ne_pars_extract  n (a b : Tm n) :
   ne a -> rtc Par.R a b -> (exists i, extract b = (VarTm i)) \/ extract b = Bot.
 Proof. hauto lq:on rew:off use:ne_pars_inv, prov_extract. Qed.
 
+Lemma const_pars_extract n k b :
+  rtc Par.R (Const k : Tm n) b -> extract b = Const k.
+Proof. hauto l:on use:pars_const_inv, prov_extract. Qed.
+
+Lemma compile_ne n (a : Tm n) :
+  ne a = ne (Compile.F a) /\ nf a = nf (Compile.F a).
+Proof.
+  elim : n / a => //=; sfirstorder b:on.
+Qed.
+
 Lemma join_bind_ne_contra n p (A : Tm n) B C :
   ne C ->
-  join (TBind p A B) C -> False.
+  Join.R (TBind p A B) C -> False.
 Proof.
-  move => hC [D [h0 h1]].
-  move /pars_pi_inv : h0 => [A0 [B0 [h2 [h3 h4]]]].
+  rewrite /Join.R. move => hC /=.
+  rewrite !pair_eq.
+  move => [[[D [h0 h1]] _] _].
+  have {}hC : ne (Compile.F C) by hauto lq:on use:compile_ne.
+  have {}hC : ne (Proj PL (Proj PL (Compile.F C))) by scongruence.
   have : (exists i, extract D = (VarTm i)) \/ extract D = Bot by eauto using ne_pars_extract.
+  have : extract D = Const p by eauto using const_pars_extract.
   sfirstorder.
 Qed.
 
 Lemma join_univ_ne_contra n i C :
   ne C ->
-  join (Univ i : Tm n) C -> False.
+  Join.R (Univ i : Tm n) C -> False.
 Proof.
   move => hC [D [h0 h1]].
   move /pars_univ_inv : h0 => ?.
-  have : (exists i, extract D = (VarTm i)) \/ extract D = Bot by eauto using ne_pars_extract.
+  have : (exists i, extract D = (VarTm i)) \/ exists k, extract D =(Const k)  by hauto q:on use:ne_pars_extract, compile_ne.
   sfirstorder.
 Qed.
 
@@ -336,7 +349,7 @@ Qed.
 Lemma InterpExt_Join n i I (A B : Tm n) PA PB :
   ⟦ A ⟧ i ;; I ↘ PA ->
   ⟦ B ⟧ i ;; I ↘ PB ->
-  join A B ->
+  Join.R A B ->
   PA = PB.
 Proof.
   move => h. move : B PB. elim : A PA /h.

From 78503149353745d8dd838169b9dbb6dbc3387172 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 20 Jan 2025 16:49:47 -0500
Subject: [PATCH 08/13] Add stub for par compile

---
 theories/compile.v | 15 +++++++++++++++
 theories/logrel.v  |  2 +-
 2 files changed, 16 insertions(+), 1 deletion(-)

diff --git a/theories/compile.v b/theories/compile.v
index 3eaaf42..f063b16 100644
--- a/theories/compile.v
+++ b/theories/compile.v
@@ -66,6 +66,18 @@ Module Join.
   Lemma UnivInj n i j : R (Univ i : Tm n) (Univ j) -> i = j.
   Proof. hauto l:on use:join_univ_inj. Qed.
 
+  Lemma transitive n (a b c : Tm n) :
+    R a b -> R b c -> R a c.
+  Proof. hauto l:on use:join_transitive unfold:R. Qed.
+
+  Lemma symmetric n (a b : Tm n) :
+    R a b -> R b a.
+  Proof. hauto l:on use:join_symmetric. Qed.
+
+  Lemma reflexive n (a : Tm n) :
+    R a a.
+  Proof. hauto l:on use:join_refl. Qed.
+
 End Join.
 
 Module Equiv.
@@ -135,3 +147,6 @@ Module EquivJoin.
     - hauto l:on.
   Qed.
 End EquivJoin.
+
+Module CompilePar.
+End CompilePar.
diff --git a/theories/logrel.v b/theories/logrel.v
index a421cda..c822d06 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -344,7 +344,7 @@ Proof.
   sfirstorder.
 Qed.
 
-#[export]Hint Resolve join_univ_ne_contra join_bind_ne_contra join_univ_pi_contra join_symmetric join_transitive : join.
+#[export]Hint Resolve join_univ_ne_contra join_bind_ne_contra join_univ_pi_contra Join.symmetric Join.transitive : join.
 
 Lemma InterpExt_Join n i I (A B : Tm n) PA PB :
   ⟦ A ⟧ i ;; I ↘ PA ->

From 36427daa61ba60fffd0308e1ec452c938ac4be43 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 20 Jan 2025 20:11:51 -0500
Subject: [PATCH 09/13] Refactor the equational theory to use the compile
 function

---
 theories/compile.v | 33 +++++++++++++++++++++++++++--
 theories/fp_red.v  | 10 ---------
 theories/logrel.v  | 52 ++++++++++++++++++++++++++++++++--------------
 3 files changed, 67 insertions(+), 28 deletions(-)

diff --git a/theories/compile.v b/theories/compile.v
index f063b16..05bc9a8 100644
--- a/theories/compile.v
+++ b/theories/compile.v
@@ -1,6 +1,7 @@
 Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red.
 Require Import ssreflect ssrbool.
 From Hammer Require Import Tactics.
+From stdpp Require Import relations (rtc(..)).
 
 Module Compile.
   Fixpoint F {n} (a : Tm n) : Tm n :=
@@ -78,6 +79,16 @@ Module Join.
     R a a.
   Proof. hauto l:on use:join_refl. Qed.
 
+  Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+    R a b -> R (subst_Tm ρ a) (subst_Tm ρ b).
+  Proof.
+    rewrite /R.
+    rewrite /join.
+    move => [C [h0 h1]].
+    repeat (rewrite <- Compile.morphing with (ρ0 := funcomp Compile.F ρ); last by reflexivity).
+    hauto lq:on use:join_substing.
+  Qed.
+
 End Join.
 
 Module Equiv.
@@ -148,5 +159,23 @@ Module EquivJoin.
   Qed.
 End EquivJoin.
 
-Module CompilePar.
-End CompilePar.
+Lemma compile_rpar n (a b : Tm n) : RPar'.R a b -> RPar'.R (Compile.F a) (Compile.F b).
+Proof.
+  move => h. elim : n a b /h.
+  - move => n  a0 a1 b0 b1 ha iha hb ihb /=.
+    rewrite -Compile.substing.
+    apply RPar'.AppAbs => //.
+  - hauto q:on use:RPar'.ProjPair'.
+  - qauto ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+  - hauto lq:on ctrs:RPar'.R.
+Qed.
+
+Lemma compile_rpars n (a b : Tm n) : rtc RPar'.R a b -> rtc RPar'.R (Compile.F a) (Compile.F b).
+Proof. induction 1; hauto lq:on ctrs:rtc use:compile_rpar. Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 6932f84..e9b5591 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2394,16 +2394,6 @@ Proof.
   sfirstorder unfold:join.
 Qed.
 
-Lemma join_univ_pi_contra n p (A : Tm n) B i :
-  join (TBind p A B) (Univ i) -> False.
-Proof.
-  rewrite /join.
-  move => [c [h0 h1]].
-  move /pars_univ_inv : h1.
-  move /pars_pi_inv : h0.
-  hauto l:on.
-Qed.
-
 Lemma join_substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
   join a b ->
   join (subst_Tm ρ a) (subst_Tm ρ b).
diff --git a/theories/logrel.v b/theories/logrel.v
index c822d06..b9c854d 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -262,8 +262,12 @@ Lemma RPar's_Pars n (A B : Tm n) :
 Proof. hauto lq:on use:RPar'_Par, rtc_subrel. Qed.
 
 Lemma RPar's_join n (A B : Tm n) :
-  rtc RPar'.R A B -> join A B.
-Proof. hauto lq:on ctrs:rtc use:RPar's_Pars. Qed.
+  rtc RPar'.R A B -> Join.R A B.
+Proof.
+  rewrite /Join.R => h.
+  have {}h : rtc RPar'.R (Compile.F A) (Compile.F B) by eauto using compile_rpars.
+  rewrite /join. eauto using RPar's_Pars, rtc_refl.
+Qed.
 
 Lemma bindspace_iff n p (PA : Tm n -> Prop) PF PF0 b  :
   (forall (a : Tm n) (PB PB0 : Tm n -> Prop), PF a PB -> PF0 a PB0 -> PB = PB0) ->
@@ -320,6 +324,19 @@ Proof.
   elim : n / a => //=; sfirstorder b:on.
 Qed.
 
+Lemma join_univ_pi_contra n p (A : Tm n) B i :
+  Join.R (TBind p A B) (Univ i) -> False.
+Proof.
+  rewrite /Join.R /=.
+  rewrite !pair_eq.
+  move => [[h _ ]_ ].
+  move : h => [C [h0 h1]].
+  have ?  : extract C = Const p by hauto l:on use:pars_const_inv.
+  have h : prov (Univ i : Tm n) (Proj PL (Proj PL (Univ i))) by hauto lq:on ctrs:prov.
+  have {h} : prov (Univ i) C by eauto using prov_pars.
+  move /prov_extract=>/=. congruence.
+Qed.
+
 Lemma join_bind_ne_contra n p (A : Tm n) B C :
   ne C ->
   Join.R (TBind p A B) C -> False.
@@ -370,9 +387,9 @@ Proof.
       move /InterpExt_Bind_inv : hPi.
       move => [PA0][PF0][hPA0][hTot0][hRes0]?. subst.
       move => hjoin.
-      have{}hr : join U (TBind p0 A0 B0) by auto using RPar's_join.
-      have hj : join (TBind p A B) (TBind p0 A0 B0) by eauto using join_transitive.
-      have {hj} : p0 = p /\ join A A0 /\ join B B0 by hauto l:on use:join_pi_inj.
+      have{}hr : Join.R U (TBind p0 A0 B0) by auto using RPar's_join.
+      have hj : Join.R (TBind p A B) (TBind p0 A0 B0) by eauto using Join.transitive.
+      have {hj} : p0 = p /\ Join.R A A0 /\ Join.R B B0 by hauto l:on use:Join.BindInj.
       move => [? [h0 h1]]. subst.
       have ? : PA0 = PA by hauto l:on. subst.
       rewrite /ProdSpace.
@@ -381,13 +398,15 @@ Proof.
       apply bindspace_iff; eauto.
       move => a PB PB0 hPB hPB0.
       apply : ihPF; eauto.
-      by apply join_substing.
+      rewrite /Join.R.
+      rewrite -!Compile.substing.
+      hauto l:on use:join_substing.
     + move => j ?. subst.
       move => [h0 h1] h.
-      have ? : join U (Univ j) by eauto using RPar's_join.
-      have : join (TBind p A B) (Univ j) by eauto using join_transitive.
+      have ? : Join.R U (Univ j) by eauto using RPar's_join.
+      have : Join.R (TBind p A B) (Univ j) by eauto using Join.transitive.
       move => ?. exfalso.
-      eauto using join_univ_pi_contra.
+      hauto l: on use: join_univ_pi_contra.
     + move => A0 ? ? [/RPar's_join ?]. subst.
       move => _ ?. exfalso. eauto with join.
   - move => j ? B PB /InterpExtInv.
@@ -397,19 +416,19 @@ Proof.
       move /RPar's_join => *.
       exfalso. eauto with join.
     + move => m _ [/RPar's_join h0 + h1].
-      have /join_univ_inj {h0 h1} ?  : join (Univ j : Tm n) (Univ m) by eauto using join_transitive.
+      have /join_univ_inj {h0 h1} ?  : Join.R (Univ j : Tm n) (Univ m) by eauto using Join.transitive.
       subst.
       move /InterpExt_Univ_inv. firstorder.
     + move => A ? ? [/RPar's_join] *. subst. exfalso. eauto with join.
   - move => A A0 PA h.
-    have /join_symmetric {}h : join A A0 by hauto lq:on ctrs:rtc use:RPar'_Par, relations.rtc_once.
-    eauto using join_transitive.
+    have /Join.symmetric {}h : Join.R A A0 by hauto lq:on ctrs:rtc use:RPar's_join, relations.rtc_once.
+    eauto with join.
 Qed.
 
 Lemma InterpUniv_Join n i (A B : Tm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
-  join A B ->
+  Join.R A B ->
   PA = PB.
 Proof. hauto l:on use:InterpExt_Join rew:db:InterpUniv. Qed.
 
@@ -442,7 +461,7 @@ Proof. hauto use:InterpExt_Functional rew:db:InterpUniv. Qed.
 Lemma InterpUniv_Join' n i j (A B : Tm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ j ↘ PB ->
-  join A B ->
+  Join.R A B ->
   PA = PB.
 Proof.
   have [? ?] : i <= max i j /\ j <= max i j by lia.
@@ -749,12 +768,12 @@ Qed.
 Lemma ST_Conv n Γ (a : Tm n) A B i :
   Γ ⊨ a ∈ A ->
   Γ ⊨ B ∈ Univ i ->
-  join A B ->
+  Join.R A B ->
   Γ ⊨ a ∈ B.
 Proof.
   move => ha /SemWt_Univ h h0.
   move => ρ hρ.
-  have {}h0 : join (subst_Tm ρ A) (subst_Tm ρ B) by eauto using join_substing.
+  have {}h0 : Join.R (subst_Tm ρ A) (subst_Tm ρ B) by eauto using Join.substing.
   move /ha : (hρ){ha} => [m [PA [h1 h2]]].
   move /h : (hρ){h} => [S hS].
   have ? : PA = S by eauto using InterpUniv_Join'. subst.
@@ -909,6 +928,7 @@ Fixpoint size_tm {n} (a : Tm n) :=
   | Proj p a => 1 + size_tm a
   | Pair a b => 1 + Nat.add (size_tm a) (size_tm b)
   | Bot => 1
+  | Const _ => 1
   | Univ _ => 1
   end.
 

From 1f7460fd11d9dc166856288ab771c8b285097c17 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 20 Jan 2025 20:42:40 -0500
Subject: [PATCH 10/13] Add new syntax for booleans

---
 syntax.sig                   |  4 ++
 theories/Autosubst2/syntax.v | 86 +++++++++++++++++++++++++++++++++++-
 theories/compile.v           |  7 ++-
 3 files changed, 95 insertions(+), 2 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index 3101cab..9e04431 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -2,6 +2,7 @@ nat  : Type
 Tm(VarTm) : Type
 PTag : Type
 TTag : Type
+bool : Type
 
 PL : PTag
 PR : PTag
@@ -15,3 +16,6 @@ TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm
 Const : TTag -> Tm
 Univ : nat -> Tm
 Bot : Tm
+BVal : bool -> Tm
+Bool : Tm
+If : Tm -> Tm -> Tm -> Tm
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index a5cb002..5b06281 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -42,7 +42,10 @@ Inductive Tm (n_Tm : nat) : Type :=
   | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
   | Const : TTag -> Tm n_Tm
   | Univ : nat -> Tm n_Tm
-  | Bot : Tm n_Tm.
+  | Bot : Tm n_Tm
+  | BVal : bool -> Tm n_Tm
+  | Bool : Tm n_Tm
+  | If : Tm n_Tm -> Tm n_Tm -> Tm n_Tm -> Tm n_Tm.
 
 Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)}
   (H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0.
@@ -101,6 +104,27 @@ Proof.
 exact (eq_refl).
 Qed.
 
+Lemma congr_BVal {m_Tm : nat} {s0 : bool} {t0 : bool} (H0 : s0 = t0) :
+  BVal m_Tm s0 = BVal m_Tm t0.
+Proof.
+exact (eq_trans eq_refl (ap (fun x => BVal m_Tm x) H0)).
+Qed.
+
+Lemma congr_Bool {m_Tm : nat} : Bool m_Tm = Bool m_Tm.
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma congr_If {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {s2 : Tm m_Tm}
+  {t0 : Tm m_Tm} {t1 : Tm m_Tm} {t2 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1)
+  (H2 : s2 = t2) : If m_Tm s0 s1 s2 = If m_Tm t0 t1 t2.
+Proof.
+exact (eq_trans
+         (eq_trans (eq_trans eq_refl (ap (fun x => If m_Tm x s1 s2) H0))
+            (ap (fun x => If m_Tm t0 x s2) H1))
+         (ap (fun x => If m_Tm t0 t1 x) H2)).
+Qed.
+
 Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) :
   fin (S m) -> fin (S n).
 Proof.
@@ -126,6 +150,10 @@ Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
   | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
   | Bot _ => Bot n_Tm
+  | BVal _ s0 => BVal n_Tm s0
+  | Bool _ => Bool n_Tm
+  | If _ s0 s1 s2 =>
+      If n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) (ren_Tm xi_Tm s2)
   end.
 
 Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) :
@@ -154,6 +182,11 @@ Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
   | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
   | Bot _ => Bot n_Tm
+  | BVal _ s0 => BVal n_Tm s0
+  | Bool _ => Bool n_Tm
+  | If _ s0 s1 s2 =>
+      If n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
+        (subst_Tm sigma_Tm s2)
   end.
 
 Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm)
@@ -195,6 +228,11 @@ subst_Tm sigma_Tm s = s :=
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
+        (idSubst_Tm sigma_Tm Eq_Tm s2)
   end.
 
 Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n)
@@ -240,6 +278,11 @@ Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0)
+        (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s2)
   end.
 
 Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm)
@@ -286,6 +329,11 @@ Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (ext_Tm sigma_Tm tau_Tm Eq_Tm s0)
+        (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) (ext_Tm sigma_Tm tau_Tm Eq_Tm s2)
   end.
 
 Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
@@ -332,6 +380,12 @@ Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0)
+        (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
+        (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s2)
   end.
 
 Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat}
@@ -389,6 +443,12 @@ Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0)
+        (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
+        (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s2)
   end.
 
 Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@@ -467,6 +527,12 @@ ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0)
+        (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
+        (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s2)
   end.
 
 Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@@ -546,6 +612,12 @@ subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0)
+        (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
+        (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s2)
   end.
 
 Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
@@ -663,6 +735,12 @@ Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
   | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
   | Bot _ => congr_Bot
+  | BVal _ s0 => congr_BVal (eq_refl s0)
+  | Bool _ => congr_Bool
+  | If _ s0 s1 s2 =>
+      congr_If (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0)
+        (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
+        (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s2)
   end.
 
 Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
@@ -861,6 +939,12 @@ Core.
 
 Arguments VarTm {n_Tm}.
 
+Arguments If {n_Tm}.
+
+Arguments Bool {n_Tm}.
+
+Arguments BVal {n_Tm}.
+
 Arguments Bot {n_Tm}.
 
 Arguments Univ {n_Tm}.
diff --git a/theories/compile.v b/theories/compile.v
index 05bc9a8..e481d71 100644
--- a/theories/compile.v
+++ b/theories/compile.v
@@ -15,11 +15,14 @@ Module Compile.
     | Pair a b => Pair (F a) (F b)
     | Proj t a => Proj t (F a)
     | Bot => Bot
+    | If a b c => App (App (F a) (F b)) (F c)
+    | BVal b => if b then (Abs (Abs (VarTm (shift var_zero)))) else (Abs (Abs (VarTm var_zero)))
+    | Bool => Bool
     end.
 
   Lemma renaming n m (a : Tm n) (ξ : fin n -> fin m) :
     F (ren_Tm ξ a)= ren_Tm ξ (F a).
-  Proof. move : m ξ. elim : n  / a => //=; scongruence. Qed.
+  Proof. move : m ξ. elim : n  / a => //=; hauto lq:on. Qed.
 
   #[local]Hint Rewrite Compile.renaming : compile.
 
@@ -33,6 +36,8 @@ Module Compile.
     - hauto lq:on rew:off.
     - hauto lq:on.
     - hauto lq:on inv:option rew:db:compile unfold:funcomp.
+    - hauto lq:on rew:off.
+    - hauto lq:on rew:off.
   Qed.
 
   Lemma substing n b (a : Tm (S n)) :

From 255bd4acbf81260957643466a8edc6f70129ba70 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 24 Jan 2025 14:52:35 -0700
Subject: [PATCH 11/13] Rename the term constructors

---
 syntax.sig                   |   12 +-
 theories/Autosubst2/syntax.v |  855 +++++++++++++++++++++++--
 theories/diagram.txt         |   22 +
 theories/fp_red.v            | 1132 +++++++++++++++-------------------
 4 files changed, 1323 insertions(+), 698 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index 9e04431..e63c3b4 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -1,4 +1,5 @@
 nat  : Type
+PTm(VarPTm) : Type
 Tm(VarTm) : Type
 PTag : Type
 TTag : Type
@@ -8,14 +9,21 @@ PL : PTag
 PR : PTag
 TPi : TTag
 TSig : TTag
+
+PAbs : (bind PTm in PTm) -> PTm
+PApp : PTm -> PTm -> PTm
+PPair : PTm -> PTm -> PTm
+PProj : PTag -> PTm -> PTm
+PConst : TTag -> PTm
+PUniv : nat -> PTm
+PBot : PTm
+
 Abs : (bind Tm in Tm) -> Tm
 App : Tm -> Tm -> Tm
 Pair : Tm -> Tm -> Tm
 Proj : PTag -> Tm -> Tm
 TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm
-Const : TTag -> Tm
 Univ : nat -> Tm
-Bot : Tm
 BVal : bool -> Tm
 Bool : Tm
 If : Tm -> Tm -> Tm -> Tm
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index 5b06281..75ef645 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -33,6 +33,674 @@ Proof.
 exact (eq_refl).
 Qed.
 
+Inductive PTm (n_PTm : nat) : Type :=
+  | VarPTm : fin n_PTm -> PTm n_PTm
+  | PAbs : PTm (S n_PTm) -> PTm n_PTm
+  | PApp : PTm n_PTm -> PTm n_PTm -> PTm n_PTm
+  | PPair : PTm n_PTm -> PTm n_PTm -> PTm n_PTm
+  | PProj : PTag -> PTm n_PTm -> PTm n_PTm
+  | PConst : TTag -> PTm n_PTm
+  | PUniv : nat -> PTm n_PTm
+  | PBot : PTm n_PTm.
+
+Lemma congr_PAbs {m_PTm : nat} {s0 : PTm (S m_PTm)} {t0 : PTm (S m_PTm)}
+  (H0 : s0 = t0) : PAbs m_PTm s0 = PAbs m_PTm t0.
+Proof.
+exact (eq_trans eq_refl (ap (fun x => PAbs m_PTm x) H0)).
+Qed.
+
+Lemma congr_PApp {m_PTm : nat} {s0 : PTm m_PTm} {s1 : PTm m_PTm}
+  {t0 : PTm m_PTm} {t1 : PTm m_PTm} (H0 : s0 = t0) (H1 : s1 = t1) :
+  PApp m_PTm s0 s1 = PApp m_PTm t0 t1.
+Proof.
+exact (eq_trans (eq_trans eq_refl (ap (fun x => PApp m_PTm x s1) H0))
+         (ap (fun x => PApp m_PTm t0 x) H1)).
+Qed.
+
+Lemma congr_PPair {m_PTm : nat} {s0 : PTm m_PTm} {s1 : PTm m_PTm}
+  {t0 : PTm m_PTm} {t1 : PTm m_PTm} (H0 : s0 = t0) (H1 : s1 = t1) :
+  PPair m_PTm s0 s1 = PPair m_PTm t0 t1.
+Proof.
+exact (eq_trans (eq_trans eq_refl (ap (fun x => PPair m_PTm x s1) H0))
+         (ap (fun x => PPair m_PTm t0 x) H1)).
+Qed.
+
+Lemma congr_PProj {m_PTm : nat} {s0 : PTag} {s1 : PTm m_PTm} {t0 : PTag}
+  {t1 : PTm m_PTm} (H0 : s0 = t0) (H1 : s1 = t1) :
+  PProj m_PTm s0 s1 = PProj m_PTm t0 t1.
+Proof.
+exact (eq_trans (eq_trans eq_refl (ap (fun x => PProj m_PTm x s1) H0))
+         (ap (fun x => PProj m_PTm t0 x) H1)).
+Qed.
+
+Lemma congr_PConst {m_PTm : nat} {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
+  PConst m_PTm s0 = PConst m_PTm t0.
+Proof.
+exact (eq_trans eq_refl (ap (fun x => PConst m_PTm x) H0)).
+Qed.
+
+Lemma congr_PUniv {m_PTm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
+  PUniv m_PTm s0 = PUniv m_PTm t0.
+Proof.
+exact (eq_trans eq_refl (ap (fun x => PUniv m_PTm x) H0)).
+Qed.
+
+Lemma congr_PBot {m_PTm : nat} : PBot m_PTm = PBot m_PTm.
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma upRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n) :
+  fin (S m) -> fin (S n).
+Proof.
+exact (up_ren xi).
+Defined.
+
+Lemma upRen_list_PTm_PTm (p : nat) {m : nat} {n : nat} (xi : fin m -> fin n)
+  : fin (plus p m) -> fin (plus p n).
+Proof.
+exact (upRen_p p xi).
+Defined.
+
+Fixpoint ren_PTm {m_PTm : nat} {n_PTm : nat}
+(xi_PTm : fin m_PTm -> fin n_PTm) (s : PTm m_PTm) {struct s} : PTm n_PTm :=
+  match s with
+  | VarPTm _ s0 => VarPTm n_PTm (xi_PTm s0)
+  | PAbs _ s0 => PAbs n_PTm (ren_PTm (upRen_PTm_PTm xi_PTm) s0)
+  | PApp _ s0 s1 => PApp n_PTm (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
+  | PPair _ s0 s1 => PPair n_PTm (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
+  | PProj _ s0 s1 => PProj n_PTm s0 (ren_PTm xi_PTm s1)
+  | PConst _ s0 => PConst n_PTm s0
+  | PUniv _ s0 => PUniv n_PTm s0
+  | PBot _ => PBot n_PTm
+  end.
+
+Lemma up_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm) :
+  fin (S m) -> PTm (S n_PTm).
+Proof.
+exact (scons (VarPTm (S n_PTm) var_zero) (funcomp (ren_PTm shift) sigma)).
+Defined.
+
+Lemma up_list_PTm_PTm (p : nat) {m : nat} {n_PTm : nat}
+  (sigma : fin m -> PTm n_PTm) : fin (plus p m) -> PTm (plus p n_PTm).
+Proof.
+exact (scons_p p (funcomp (VarPTm (plus p n_PTm)) (zero_p p))
+         (funcomp (ren_PTm (shift_p p)) sigma)).
+Defined.
+
+Fixpoint subst_PTm {m_PTm : nat} {n_PTm : nat}
+(sigma_PTm : fin m_PTm -> PTm n_PTm) (s : PTm m_PTm) {struct s} : PTm n_PTm
+:=
+  match s with
+  | VarPTm _ s0 => sigma_PTm s0
+  | PAbs _ s0 => PAbs n_PTm (subst_PTm (up_PTm_PTm sigma_PTm) s0)
+  | PApp _ s0 s1 =>
+      PApp n_PTm (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
+  | PPair _ s0 s1 =>
+      PPair n_PTm (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
+  | PProj _ s0 s1 => PProj n_PTm s0 (subst_PTm sigma_PTm s1)
+  | PConst _ s0 => PConst n_PTm s0
+  | PUniv _ s0 => PUniv n_PTm s0
+  | PBot _ => PBot n_PTm
+  end.
+
+Lemma upId_PTm_PTm {m_PTm : nat} (sigma : fin m_PTm -> PTm m_PTm)
+  (Eq : forall x, sigma x = VarPTm m_PTm x) :
+  forall x, up_PTm_PTm sigma x = VarPTm (S m_PTm) x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma upId_list_PTm_PTm {p : nat} {m_PTm : nat}
+  (sigma : fin m_PTm -> PTm m_PTm) (Eq : forall x, sigma x = VarPTm m_PTm x)
+  : forall x, up_list_PTm_PTm p sigma x = VarPTm (plus p m_PTm) x.
+Proof.
+exact (fun n =>
+       scons_p_eta (VarPTm (plus p m_PTm))
+         (fun n => ap (ren_PTm (shift_p p)) (Eq n)) (fun n => eq_refl)).
+Qed.
+
+Fixpoint idSubst_PTm {m_PTm : nat} (sigma_PTm : fin m_PTm -> PTm m_PTm)
+(Eq_PTm : forall x, sigma_PTm x = VarPTm m_PTm x) (s : PTm m_PTm) {struct s}
+   : subst_PTm sigma_PTm s = s :=
+  match s with
+  | VarPTm _ s0 => Eq_PTm s0
+  | PAbs _ s0 =>
+      congr_PAbs
+        (idSubst_PTm (up_PTm_PTm sigma_PTm) (upId_PTm_PTm _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (idSubst_PTm sigma_PTm Eq_PTm s0)
+        (idSubst_PTm sigma_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (idSubst_PTm sigma_PTm Eq_PTm s0)
+        (idSubst_PTm sigma_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma upExtRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n)
+  (zeta : fin m -> fin n) (Eq : forall x, xi x = zeta x) :
+  forall x, upRen_PTm_PTm xi x = upRen_PTm_PTm zeta x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n => ap shift (Eq fin_n)
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma upExtRen_list_PTm_PTm {p : nat} {m : nat} {n : nat}
+  (xi : fin m -> fin n) (zeta : fin m -> fin n)
+  (Eq : forall x, xi x = zeta x) :
+  forall x, upRen_list_PTm_PTm p xi x = upRen_list_PTm_PTm p zeta x.
+Proof.
+exact (fun n =>
+       scons_p_congr (fun n => eq_refl) (fun n => ap (shift_p p) (Eq n))).
+Qed.
+
+Fixpoint extRen_PTm {m_PTm : nat} {n_PTm : nat}
+(xi_PTm : fin m_PTm -> fin n_PTm) (zeta_PTm : fin m_PTm -> fin n_PTm)
+(Eq_PTm : forall x, xi_PTm x = zeta_PTm x) (s : PTm m_PTm) {struct s} :
+ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
+  match s with
+  | VarPTm _ s0 => ap (VarPTm n_PTm) (Eq_PTm s0)
+  | PAbs _ s0 =>
+      congr_PAbs
+        (extRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
+           (upExtRen_PTm_PTm _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0)
+        (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0)
+        (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0) (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma upExt_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm)
+  (tau : fin m -> PTm n_PTm) (Eq : forall x, sigma x = tau x) :
+  forall x, up_PTm_PTm sigma x = up_PTm_PTm tau x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma upExt_list_PTm_PTm {p : nat} {m : nat} {n_PTm : nat}
+  (sigma : fin m -> PTm n_PTm) (tau : fin m -> PTm n_PTm)
+  (Eq : forall x, sigma x = tau x) :
+  forall x, up_list_PTm_PTm p sigma x = up_list_PTm_PTm p tau x.
+Proof.
+exact (fun n =>
+       scons_p_congr (fun n => eq_refl)
+         (fun n => ap (ren_PTm (shift_p p)) (Eq n))).
+Qed.
+
+Fixpoint ext_PTm {m_PTm : nat} {n_PTm : nat}
+(sigma_PTm : fin m_PTm -> PTm n_PTm) (tau_PTm : fin m_PTm -> PTm n_PTm)
+(Eq_PTm : forall x, sigma_PTm x = tau_PTm x) (s : PTm m_PTm) {struct s} :
+subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
+  match s with
+  | VarPTm _ s0 => Eq_PTm s0
+  | PAbs _ s0 =>
+      congr_PAbs
+        (ext_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
+           (upExt_PTm_PTm _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (ext_PTm sigma_PTm tau_PTm Eq_PTm s0)
+        (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (ext_PTm sigma_PTm tau_PTm Eq_PTm s0)
+        (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0) (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma up_ren_ren_PTm_PTm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
+  (zeta : fin l -> fin m) (rho : fin k -> fin m)
+  (Eq : forall x, funcomp zeta xi x = rho x) :
+  forall x,
+  funcomp (upRen_PTm_PTm zeta) (upRen_PTm_PTm xi) x = upRen_PTm_PTm rho x.
+Proof.
+exact (up_ren_ren xi zeta rho Eq).
+Qed.
+
+Lemma up_ren_ren_list_PTm_PTm {p : nat} {k : nat} {l : nat} {m : nat}
+  (xi : fin k -> fin l) (zeta : fin l -> fin m) (rho : fin k -> fin m)
+  (Eq : forall x, funcomp zeta xi x = rho x) :
+  forall x,
+  funcomp (upRen_list_PTm_PTm p zeta) (upRen_list_PTm_PTm p xi) x =
+  upRen_list_PTm_PTm p rho x.
+Proof.
+exact (up_ren_ren_p Eq).
+Qed.
+
+Fixpoint compRenRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+(xi_PTm : fin m_PTm -> fin k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
+(rho_PTm : fin m_PTm -> fin l_PTm)
+(Eq_PTm : forall x, funcomp zeta_PTm xi_PTm x = rho_PTm x) (s : PTm m_PTm)
+{struct s} : ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm rho_PTm s :=
+  match s with
+  | VarPTm _ s0 => ap (VarPTm l_PTm) (Eq_PTm s0)
+  | PAbs _ s0 =>
+      congr_PAbs
+        (compRenRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
+           (upRen_PTm_PTm rho_PTm) (up_ren_ren _ _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0)
+        (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0)
+        (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0)
+        (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma up_ren_subst_PTm_PTm {k : nat} {l : nat} {m_PTm : nat}
+  (xi : fin k -> fin l) (tau : fin l -> PTm m_PTm)
+  (theta : fin k -> PTm m_PTm) (Eq : forall x, funcomp tau xi x = theta x) :
+  forall x,
+  funcomp (up_PTm_PTm tau) (upRen_PTm_PTm xi) x = up_PTm_PTm theta x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma up_ren_subst_list_PTm_PTm {p : nat} {k : nat} {l : nat} {m_PTm : nat}
+  (xi : fin k -> fin l) (tau : fin l -> PTm m_PTm)
+  (theta : fin k -> PTm m_PTm) (Eq : forall x, funcomp tau xi x = theta x) :
+  forall x,
+  funcomp (up_list_PTm_PTm p tau) (upRen_list_PTm_PTm p xi) x =
+  up_list_PTm_PTm p theta x.
+Proof.
+exact (fun n =>
+       eq_trans (scons_p_comp' _ _ _ n)
+         (scons_p_congr (fun z => scons_p_head' _ _ z)
+            (fun z =>
+             eq_trans (scons_p_tail' _ _ (xi z))
+               (ap (ren_PTm (shift_p p)) (Eq z))))).
+Qed.
+
+Fixpoint compRenSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+(xi_PTm : fin m_PTm -> fin k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
+(theta_PTm : fin m_PTm -> PTm l_PTm)
+(Eq_PTm : forall x, funcomp tau_PTm xi_PTm x = theta_PTm x) (s : PTm m_PTm)
+{struct s} : subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm theta_PTm s :=
+  match s with
+  | VarPTm _ s0 => Eq_PTm s0
+  | PAbs _ s0 =>
+      congr_PAbs
+        (compRenSubst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm tau_PTm)
+           (up_PTm_PTm theta_PTm) (up_ren_subst_PTm_PTm _ _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0)
+        (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0)
+        (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0)
+        (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma up_subst_ren_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
+  (sigma : fin k -> PTm l_PTm) (zeta_PTm : fin l_PTm -> fin m_PTm)
+  (theta : fin k -> PTm m_PTm)
+  (Eq : forall x, funcomp (ren_PTm zeta_PTm) sigma x = theta x) :
+  forall x,
+  funcomp (ren_PTm (upRen_PTm_PTm zeta_PTm)) (up_PTm_PTm sigma) x =
+  up_PTm_PTm theta x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n =>
+           eq_trans
+             (compRenRen_PTm shift (upRen_PTm_PTm zeta_PTm)
+                (funcomp shift zeta_PTm) (fun x => eq_refl) (sigma fin_n))
+             (eq_trans
+                (eq_sym
+                   (compRenRen_PTm zeta_PTm shift (funcomp shift zeta_PTm)
+                      (fun x => eq_refl) (sigma fin_n)))
+                (ap (ren_PTm shift) (Eq fin_n)))
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma up_subst_ren_list_PTm_PTm {p : nat} {k : nat} {l_PTm : nat}
+  {m_PTm : nat} (sigma : fin k -> PTm l_PTm)
+  (zeta_PTm : fin l_PTm -> fin m_PTm) (theta : fin k -> PTm m_PTm)
+  (Eq : forall x, funcomp (ren_PTm zeta_PTm) sigma x = theta x) :
+  forall x,
+  funcomp (ren_PTm (upRen_list_PTm_PTm p zeta_PTm)) (up_list_PTm_PTm p sigma)
+    x = up_list_PTm_PTm p theta x.
+Proof.
+exact (fun n =>
+       eq_trans (scons_p_comp' _ _ _ n)
+         (scons_p_congr
+            (fun x => ap (VarPTm (plus p m_PTm)) (scons_p_head' _ _ x))
+            (fun n =>
+             eq_trans
+               (compRenRen_PTm (shift_p p) (upRen_list_PTm_PTm p zeta_PTm)
+                  (funcomp (shift_p p) zeta_PTm)
+                  (fun x => scons_p_tail' _ _ x) (sigma n))
+               (eq_trans
+                  (eq_sym
+                     (compRenRen_PTm zeta_PTm (shift_p p)
+                        (funcomp (shift_p p) zeta_PTm) (fun x => eq_refl)
+                        (sigma n))) (ap (ren_PTm (shift_p p)) (Eq n)))))).
+Qed.
+
+Fixpoint compSubstRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+(sigma_PTm : fin m_PTm -> PTm k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
+(theta_PTm : fin m_PTm -> PTm l_PTm)
+(Eq_PTm : forall x, funcomp (ren_PTm zeta_PTm) sigma_PTm x = theta_PTm x)
+(s : PTm m_PTm) {struct s} :
+ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
+  match s with
+  | VarPTm _ s0 => Eq_PTm s0
+  | PAbs _ s0 =>
+      congr_PAbs
+        (compSubstRen_PTm (up_PTm_PTm sigma_PTm) (upRen_PTm_PTm zeta_PTm)
+           (up_PTm_PTm theta_PTm) (up_subst_ren_PTm_PTm _ _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0)
+        (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0)
+        (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0)
+        (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma up_subst_subst_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
+  (sigma : fin k -> PTm l_PTm) (tau_PTm : fin l_PTm -> PTm m_PTm)
+  (theta : fin k -> PTm m_PTm)
+  (Eq : forall x, funcomp (subst_PTm tau_PTm) sigma x = theta x) :
+  forall x,
+  funcomp (subst_PTm (up_PTm_PTm tau_PTm)) (up_PTm_PTm sigma) x =
+  up_PTm_PTm theta x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n =>
+           eq_trans
+             (compRenSubst_PTm shift (up_PTm_PTm tau_PTm)
+                (funcomp (up_PTm_PTm tau_PTm) shift) (fun x => eq_refl)
+                (sigma fin_n))
+             (eq_trans
+                (eq_sym
+                   (compSubstRen_PTm tau_PTm shift
+                      (funcomp (ren_PTm shift) tau_PTm) (fun x => eq_refl)
+                      (sigma fin_n))) (ap (ren_PTm shift) (Eq fin_n)))
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma up_subst_subst_list_PTm_PTm {p : nat} {k : nat} {l_PTm : nat}
+  {m_PTm : nat} (sigma : fin k -> PTm l_PTm)
+  (tau_PTm : fin l_PTm -> PTm m_PTm) (theta : fin k -> PTm m_PTm)
+  (Eq : forall x, funcomp (subst_PTm tau_PTm) sigma x = theta x) :
+  forall x,
+  funcomp (subst_PTm (up_list_PTm_PTm p tau_PTm)) (up_list_PTm_PTm p sigma) x =
+  up_list_PTm_PTm p theta x.
+Proof.
+exact (fun n =>
+       eq_trans
+         (scons_p_comp' (funcomp (VarPTm (plus p l_PTm)) (zero_p p)) _ _ n)
+         (scons_p_congr
+            (fun x => scons_p_head' _ (fun z => ren_PTm (shift_p p) _) x)
+            (fun n =>
+             eq_trans
+               (compRenSubst_PTm (shift_p p) (up_list_PTm_PTm p tau_PTm)
+                  (funcomp (up_list_PTm_PTm p tau_PTm) (shift_p p))
+                  (fun x => eq_refl) (sigma n))
+               (eq_trans
+                  (eq_sym
+                     (compSubstRen_PTm tau_PTm (shift_p p) _
+                        (fun x => eq_sym (scons_p_tail' _ _ x)) (sigma n)))
+                  (ap (ren_PTm (shift_p p)) (Eq n)))))).
+Qed.
+
+Fixpoint compSubstSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+(sigma_PTm : fin m_PTm -> PTm k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
+(theta_PTm : fin m_PTm -> PTm l_PTm)
+(Eq_PTm : forall x, funcomp (subst_PTm tau_PTm) sigma_PTm x = theta_PTm x)
+(s : PTm m_PTm) {struct s} :
+subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
+  match s with
+  | VarPTm _ s0 => Eq_PTm s0
+  | PAbs _ s0 =>
+      congr_PAbs
+        (compSubstSubst_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
+           (up_PTm_PTm theta_PTm) (up_subst_subst_PTm_PTm _ _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0)
+        (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0)
+        (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0)
+        (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma renRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
+  (s : PTm m_PTm) :
+  ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm (funcomp zeta_PTm xi_PTm) s.
+Proof.
+exact (compRenRen_PTm xi_PTm zeta_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma renRen'_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm) :
+  pointwise_relation _ eq (funcomp (ren_PTm zeta_PTm) (ren_PTm xi_PTm))
+    (ren_PTm (funcomp zeta_PTm xi_PTm)).
+Proof.
+exact (fun s => compRenRen_PTm xi_PTm zeta_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma renSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
+  (s : PTm m_PTm) :
+  subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm (funcomp tau_PTm xi_PTm) s.
+Proof.
+exact (compRenSubst_PTm xi_PTm tau_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma renSubst_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm) :
+  pointwise_relation _ eq (funcomp (subst_PTm tau_PTm) (ren_PTm xi_PTm))
+    (subst_PTm (funcomp tau_PTm xi_PTm)).
+Proof.
+exact (fun s => compRenSubst_PTm xi_PTm tau_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma substRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (sigma_PTm : fin m_PTm -> PTm k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
+  (s : PTm m_PTm) :
+  ren_PTm zeta_PTm (subst_PTm sigma_PTm s) =
+  subst_PTm (funcomp (ren_PTm zeta_PTm) sigma_PTm) s.
+Proof.
+exact (compSubstRen_PTm sigma_PTm zeta_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma substRen_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (sigma_PTm : fin m_PTm -> PTm k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm) :
+  pointwise_relation _ eq (funcomp (ren_PTm zeta_PTm) (subst_PTm sigma_PTm))
+    (subst_PTm (funcomp (ren_PTm zeta_PTm) sigma_PTm)).
+Proof.
+exact (fun s => compSubstRen_PTm sigma_PTm zeta_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma substSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (sigma_PTm : fin m_PTm -> PTm k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
+  (s : PTm m_PTm) :
+  subst_PTm tau_PTm (subst_PTm sigma_PTm s) =
+  subst_PTm (funcomp (subst_PTm tau_PTm) sigma_PTm) s.
+Proof.
+exact (compSubstSubst_PTm sigma_PTm tau_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma substSubst_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
+  (sigma_PTm : fin m_PTm -> PTm k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm) :
+  pointwise_relation _ eq (funcomp (subst_PTm tau_PTm) (subst_PTm sigma_PTm))
+    (subst_PTm (funcomp (subst_PTm tau_PTm) sigma_PTm)).
+Proof.
+exact (fun s => compSubstSubst_PTm sigma_PTm tau_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma rinstInst_up_PTm_PTm {m : nat} {n_PTm : nat} (xi : fin m -> fin n_PTm)
+  (sigma : fin m -> PTm n_PTm)
+  (Eq : forall x, funcomp (VarPTm n_PTm) xi x = sigma x) :
+  forall x,
+  funcomp (VarPTm (S n_PTm)) (upRen_PTm_PTm xi) x = up_PTm_PTm sigma x.
+Proof.
+exact (fun n =>
+       match n with
+       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
+       | None => eq_refl
+       end).
+Qed.
+
+Lemma rinstInst_up_list_PTm_PTm {p : nat} {m : nat} {n_PTm : nat}
+  (xi : fin m -> fin n_PTm) (sigma : fin m -> PTm n_PTm)
+  (Eq : forall x, funcomp (VarPTm n_PTm) xi x = sigma x) :
+  forall x,
+  funcomp (VarPTm (plus p n_PTm)) (upRen_list_PTm_PTm p xi) x =
+  up_list_PTm_PTm p sigma x.
+Proof.
+exact (fun n =>
+       eq_trans (scons_p_comp' _ _ (VarPTm (plus p n_PTm)) n)
+         (scons_p_congr (fun z => eq_refl)
+            (fun n => ap (ren_PTm (shift_p p)) (Eq n)))).
+Qed.
+
+Fixpoint rinst_inst_PTm {m_PTm : nat} {n_PTm : nat}
+(xi_PTm : fin m_PTm -> fin n_PTm) (sigma_PTm : fin m_PTm -> PTm n_PTm)
+(Eq_PTm : forall x, funcomp (VarPTm n_PTm) xi_PTm x = sigma_PTm x)
+(s : PTm m_PTm) {struct s} : ren_PTm xi_PTm s = subst_PTm sigma_PTm s :=
+  match s with
+  | VarPTm _ s0 => Eq_PTm s0
+  | PAbs _ s0 =>
+      congr_PAbs
+        (rinst_inst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm sigma_PTm)
+           (rinstInst_up_PTm_PTm _ _ Eq_PTm) s0)
+  | PApp _ s0 s1 =>
+      congr_PApp (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0)
+        (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
+  | PPair _ s0 s1 =>
+      congr_PPair (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0)
+        (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
+  | PProj _ s0 s1 =>
+      congr_PProj (eq_refl s0) (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
+  | PConst _ s0 => congr_PConst (eq_refl s0)
+  | PUniv _ s0 => congr_PUniv (eq_refl s0)
+  | PBot _ => congr_PBot
+  end.
+
+Lemma rinstInst'_PTm {m_PTm : nat} {n_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin n_PTm) (s : PTm m_PTm) :
+  ren_PTm xi_PTm s = subst_PTm (funcomp (VarPTm n_PTm) xi_PTm) s.
+Proof.
+exact (rinst_inst_PTm xi_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma rinstInst'_PTm_pointwise {m_PTm : nat} {n_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin n_PTm) :
+  pointwise_relation _ eq (ren_PTm xi_PTm)
+    (subst_PTm (funcomp (VarPTm n_PTm) xi_PTm)).
+Proof.
+exact (fun s => rinst_inst_PTm xi_PTm _ (fun n => eq_refl) s).
+Qed.
+
+Lemma instId'_PTm {m_PTm : nat} (s : PTm m_PTm) :
+  subst_PTm (VarPTm m_PTm) s = s.
+Proof.
+exact (idSubst_PTm (VarPTm m_PTm) (fun n => eq_refl) s).
+Qed.
+
+Lemma instId'_PTm_pointwise {m_PTm : nat} :
+  pointwise_relation _ eq (subst_PTm (VarPTm m_PTm)) id.
+Proof.
+exact (fun s => idSubst_PTm (VarPTm m_PTm) (fun n => eq_refl) s).
+Qed.
+
+Lemma rinstId'_PTm {m_PTm : nat} (s : PTm m_PTm) : ren_PTm id s = s.
+Proof.
+exact (eq_ind_r (fun t => t = s) (instId'_PTm s) (rinstInst'_PTm id s)).
+Qed.
+
+Lemma rinstId'_PTm_pointwise {m_PTm : nat} :
+  pointwise_relation _ eq (@ren_PTm m_PTm m_PTm id) id.
+Proof.
+exact (fun s =>
+       eq_ind_r (fun t => t = s) (instId'_PTm s) (rinstInst'_PTm id s)).
+Qed.
+
+Lemma varL'_PTm {m_PTm : nat} {n_PTm : nat}
+  (sigma_PTm : fin m_PTm -> PTm n_PTm) (x : fin m_PTm) :
+  subst_PTm sigma_PTm (VarPTm m_PTm x) = sigma_PTm x.
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma varL'_PTm_pointwise {m_PTm : nat} {n_PTm : nat}
+  (sigma_PTm : fin m_PTm -> PTm n_PTm) :
+  pointwise_relation _ eq (funcomp (subst_PTm sigma_PTm) (VarPTm m_PTm))
+    sigma_PTm.
+Proof.
+exact (fun x => eq_refl).
+Qed.
+
+Lemma varLRen'_PTm {m_PTm : nat} {n_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin n_PTm) (x : fin m_PTm) :
+  ren_PTm xi_PTm (VarPTm m_PTm x) = VarPTm n_PTm (xi_PTm x).
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma varLRen'_PTm_pointwise {m_PTm : nat} {n_PTm : nat}
+  (xi_PTm : fin m_PTm -> fin n_PTm) :
+  pointwise_relation _ eq (funcomp (ren_PTm xi_PTm) (VarPTm m_PTm))
+    (funcomp (VarPTm n_PTm) xi_PTm).
+Proof.
+exact (fun x => eq_refl).
+Qed.
+
 Inductive Tm (n_Tm : nat) : Type :=
   | VarTm : fin n_Tm -> Tm n_Tm
   | Abs : Tm (S n_Tm) -> Tm n_Tm
@@ -40,9 +708,7 @@ Inductive Tm (n_Tm : nat) : Type :=
   | Pair : Tm n_Tm -> Tm n_Tm -> Tm n_Tm
   | Proj : PTag -> Tm n_Tm -> Tm n_Tm
   | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
-  | Const : TTag -> Tm n_Tm
   | Univ : nat -> Tm n_Tm
-  | Bot : Tm n_Tm
   | BVal : bool -> Tm n_Tm
   | Bool : Tm n_Tm
   | If : Tm n_Tm -> Tm n_Tm -> Tm n_Tm -> Tm n_Tm.
@@ -87,23 +753,12 @@ exact (eq_trans
          (ap (fun x => TBind m_Tm t0 t1 x) H2)).
 Qed.
 
-Lemma congr_Const {m_Tm : nat} {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
-  Const m_Tm s0 = Const m_Tm t0.
-Proof.
-exact (eq_trans eq_refl (ap (fun x => Const m_Tm x) H0)).
-Qed.
-
 Lemma congr_Univ {m_Tm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
   Univ m_Tm s0 = Univ m_Tm t0.
 Proof.
 exact (eq_trans eq_refl (ap (fun x => Univ m_Tm x) H0)).
 Qed.
 
-Lemma congr_Bot {m_Tm : nat} : Bot m_Tm = Bot m_Tm.
-Proof.
-exact (eq_refl).
-Qed.
-
 Lemma congr_BVal {m_Tm : nat} {s0 : bool} {t0 : bool} (H0 : s0 = t0) :
   BVal m_Tm s0 = BVal m_Tm t0.
 Proof.
@@ -147,9 +802,7 @@ Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
   | Proj _ s0 s1 => Proj n_Tm s0 (ren_Tm xi_Tm s1)
   | TBind _ s0 s1 s2 =>
       TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
-  | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
-  | Bot _ => Bot n_Tm
   | BVal _ s0 => BVal n_Tm s0
   | Bool _ => Bool n_Tm
   | If _ s0 s1 s2 =>
@@ -179,9 +832,7 @@ Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
   | Proj _ s0 s1 => Proj n_Tm s0 (subst_Tm sigma_Tm s1)
   | TBind _ s0 s1 s2 =>
       TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
-  | Const _ s0 => Const n_Tm s0
   | Univ _ s0 => Univ n_Tm s0
-  | Bot _ => Bot n_Tm
   | BVal _ s0 => BVal n_Tm s0
   | Bool _ => Bool n_Tm
   | If _ s0 s1 s2 =>
@@ -225,9 +876,7 @@ subst_Tm sigma_Tm s = s :=
   | TBind _ s0 s1 s2 =>
       congr_TBind (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
         (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -275,9 +924,7 @@ Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
       congr_TBind (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
         (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upExtRen_Tm_Tm _ _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -326,9 +973,7 @@ Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
       congr_TBind (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
         (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm)
            s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -377,9 +1022,7 @@ Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
       congr_TBind (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
         (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -440,9 +1083,7 @@ Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -524,9 +1165,7 @@ ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
         (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -609,9 +1248,7 @@ subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -732,9 +1369,7 @@ Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
       congr_TBind (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
         (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm)
            (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2)
-  | Const _ s0 => congr_Const (eq_refl s0)
   | Univ _ s0 => congr_Univ (eq_refl s0)
-  | Bot _ => congr_Bot
   | BVal _ s0 => congr_BVal (eq_refl s0)
   | Bool _ => congr_Bool
   | If _ s0 s1 s2 =>
@@ -809,6 +1444,9 @@ Qed.
 Class Up_Tm X Y :=
     up_Tm : X -> Y.
 
+Class Up_PTm X Y :=
+    up_PTm : X -> Y.
+
 #[global]
 Instance Subst_Tm  {m_Tm n_Tm : nat}: (Subst1 _ _ _) := (@subst_Tm m_Tm n_Tm).
 
@@ -821,6 +1459,19 @@ Instance Ren_Tm  {m_Tm n_Tm : nat}: (Ren1 _ _ _) := (@ren_Tm m_Tm n_Tm).
 #[global]
 Instance VarInstance_Tm  {n_Tm : nat}: (Var _ _) := (@VarTm n_Tm).
 
+#[global]
+Instance Subst_PTm  {m_PTm n_PTm : nat}: (Subst1 _ _ _) :=
+ (@subst_PTm m_PTm n_PTm).
+
+#[global]
+Instance Up_PTm_PTm  {m n_PTm : nat}: (Up_PTm _ _) := (@up_PTm_PTm m n_PTm).
+
+#[global]
+Instance Ren_PTm  {m_PTm n_PTm : nat}: (Ren1 _ _ _) := (@ren_PTm m_PTm n_PTm).
+
+#[global]
+Instance VarInstance_PTm  {n_PTm : nat}: (Var _ _) := (@VarPTm n_PTm).
+
 Notation "[ sigma_Tm ]" := (subst_Tm sigma_Tm)
 ( at level 1, left associativity, only printing)  : fscope.
 
@@ -845,6 +1496,30 @@ Notation "x '__Tm'" := (@ids _ _ VarInstance_Tm x)
 Notation "x '__Tm'" := (VarTm x) ( at level 5, format "x __Tm")  :
 subst_scope.
 
+Notation "[ sigma_PTm ]" := (subst_PTm sigma_PTm)
+( at level 1, left associativity, only printing)  : fscope.
+
+Notation "s [ sigma_PTm ]" := (subst_PTm sigma_PTm s)
+( at level 7, left associativity, only printing)  : subst_scope.
+
+Notation "↑__PTm" := up_PTm (only printing)  : subst_scope.
+
+Notation "↑__PTm" := up_PTm_PTm (only printing)  : subst_scope.
+
+Notation "⟨ xi_PTm ⟩" := (ren_PTm xi_PTm)
+( at level 1, left associativity, only printing)  : fscope.
+
+Notation "s ⟨ xi_PTm ⟩" := (ren_PTm xi_PTm s)
+( at level 7, left associativity, only printing)  : subst_scope.
+
+Notation "'var'" := VarPTm ( at level 1, only printing)  : subst_scope.
+
+Notation "x '__PTm'" := (@ids _ _ VarInstance_PTm x)
+( at level 5, format "x __PTm", only printing)  : subst_scope.
+
+Notation "x '__PTm'" := (VarPTm x) ( at level 5, format "x __PTm")  :
+subst_scope.
+
 #[global]
 Instance subst_Tm_morphism  {m_Tm : nat} {n_Tm : nat}:
  (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
@@ -881,15 +1556,56 @@ Proof.
 exact (fun f_Tm g_Tm Eq_Tm s => extRen_Tm f_Tm g_Tm Eq_Tm s).
 Qed.
 
+#[global]
+Instance subst_PTm_morphism  {m_PTm : nat} {n_PTm : nat}:
+ (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
+    (@subst_PTm m_PTm n_PTm)).
+Proof.
+exact (fun f_PTm g_PTm Eq_PTm s t Eq_st =>
+       eq_ind s (fun t' => subst_PTm f_PTm s = subst_PTm g_PTm t')
+         (ext_PTm f_PTm g_PTm Eq_PTm s) t Eq_st).
+Qed.
+
+#[global]
+Instance subst_PTm_morphism2  {m_PTm : nat} {n_PTm : nat}:
+ (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq))
+    (@subst_PTm m_PTm n_PTm)).
+Proof.
+exact (fun f_PTm g_PTm Eq_PTm s => ext_PTm f_PTm g_PTm Eq_PTm s).
+Qed.
+
+#[global]
+Instance ren_PTm_morphism  {m_PTm : nat} {n_PTm : nat}:
+ (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
+    (@ren_PTm m_PTm n_PTm)).
+Proof.
+exact (fun f_PTm g_PTm Eq_PTm s t Eq_st =>
+       eq_ind s (fun t' => ren_PTm f_PTm s = ren_PTm g_PTm t')
+         (extRen_PTm f_PTm g_PTm Eq_PTm s) t Eq_st).
+Qed.
+
+#[global]
+Instance ren_PTm_morphism2  {m_PTm : nat} {n_PTm : nat}:
+ (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq))
+    (@ren_PTm m_PTm n_PTm)).
+Proof.
+exact (fun f_PTm g_PTm Eq_PTm s => extRen_PTm f_PTm g_PTm Eq_PTm s).
+Qed.
+
 Ltac auto_unfold := repeat
-                     unfold VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1,
-                      Up_Tm_Tm, Up_Tm, up_Tm, Subst_Tm, Subst1, subst1.
+                     unfold VarInstance_PTm, Var, ids, Ren_PTm, Ren1, ren1,
+                      Up_PTm_PTm, Up_PTm, up_PTm, Subst_PTm, Subst1, subst1,
+                      VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1, Up_Tm_Tm,
+                      Up_Tm, up_Tm, Subst_Tm, Subst1, subst1.
 
 Tactic Notation "auto_unfold" "in" "*" := repeat
-                                           unfold VarInstance_Tm, Var, ids,
-                                            Ren_Tm, Ren1, ren1, Up_Tm_Tm,
-                                            Up_Tm, up_Tm, Subst_Tm, Subst1,
-                                            subst1 in *.
+                                           unfold VarInstance_PTm, Var, ids,
+                                            Ren_PTm, Ren1, ren1, Up_PTm_PTm,
+                                            Up_PTm, up_PTm, Subst_PTm,
+                                            Subst1, subst1, VarInstance_Tm,
+                                            Var, ids, Ren_Tm, Ren1, ren1,
+                                            Up_Tm_Tm, Up_Tm, up_Tm, Subst_Tm,
+                                            Subst1, subst1 in *.
 
 Ltac asimpl' := repeat (first
                  [ progress setoid_rewrite substSubst_Tm_pointwise
@@ -900,6 +1616,14 @@ Ltac asimpl' := repeat (first
                  | progress setoid_rewrite renSubst_Tm
                  | progress setoid_rewrite renRen'_Tm_pointwise
                  | progress setoid_rewrite renRen_Tm
+                 | progress setoid_rewrite substSubst_PTm_pointwise
+                 | progress setoid_rewrite substSubst_PTm
+                 | progress setoid_rewrite substRen_PTm_pointwise
+                 | progress setoid_rewrite substRen_PTm
+                 | progress setoid_rewrite renSubst_PTm_pointwise
+                 | progress setoid_rewrite renSubst_PTm
+                 | progress setoid_rewrite renRen'_PTm_pointwise
+                 | progress setoid_rewrite renRen_PTm
                  | progress setoid_rewrite varLRen'_Tm_pointwise
                  | progress setoid_rewrite varLRen'_Tm
                  | progress setoid_rewrite varL'_Tm_pointwise
@@ -908,27 +1632,42 @@ Ltac asimpl' := repeat (first
                  | progress setoid_rewrite rinstId'_Tm
                  | progress setoid_rewrite instId'_Tm_pointwise
                  | progress setoid_rewrite instId'_Tm
+                 | progress setoid_rewrite varLRen'_PTm_pointwise
+                 | progress setoid_rewrite varLRen'_PTm
+                 | progress setoid_rewrite varL'_PTm_pointwise
+                 | progress setoid_rewrite varL'_PTm
+                 | progress setoid_rewrite rinstId'_PTm_pointwise
+                 | progress setoid_rewrite rinstId'_PTm
+                 | progress setoid_rewrite instId'_PTm_pointwise
+                 | progress setoid_rewrite instId'_PTm
                  | progress
                     unfold up_list_Tm_Tm, up_Tm_Tm, upRen_list_Tm_Tm,
-                     upRen_Tm_Tm, up_ren
-                 | progress cbn[subst_Tm ren_Tm]
+                     upRen_Tm_Tm, up_list_PTm_PTm, up_PTm_PTm,
+                     upRen_list_PTm_PTm, upRen_PTm_PTm, up_ren
+                 | progress cbn[subst_Tm ren_Tm subst_PTm ren_PTm]
                  | progress fsimpl ]).
 
 Ltac asimpl := check_no_evars;
                 repeat
-                 unfold VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1,
-                  Up_Tm_Tm, Up_Tm, up_Tm, Subst_Tm, Subst1, subst1 in *;
-                asimpl'; minimize.
+                 unfold VarInstance_PTm, Var, ids, Ren_PTm, Ren1, ren1,
+                  Up_PTm_PTm, Up_PTm, up_PTm, Subst_PTm, Subst1, subst1,
+                  VarInstance_Tm, Var, ids, Ren_Tm, Ren1, ren1, Up_Tm_Tm,
+                  Up_Tm, up_Tm, Subst_Tm, Subst1, subst1 in *; asimpl';
+                minimize.
 
 Tactic Notation "asimpl" "in" hyp(J) := revert J; asimpl; intros J.
 
 Tactic Notation "auto_case" := auto_case ltac:(asimpl; cbn; eauto).
 
 Ltac substify := auto_unfold; try setoid_rewrite rinstInst'_Tm_pointwise;
-                  try setoid_rewrite rinstInst'_Tm.
+                  try setoid_rewrite rinstInst'_Tm;
+                  try setoid_rewrite rinstInst'_PTm_pointwise;
+                  try setoid_rewrite rinstInst'_PTm.
 
 Ltac renamify := auto_unfold; try setoid_rewrite_left rinstInst'_Tm_pointwise;
-                  try setoid_rewrite_left rinstInst'_Tm.
+                  try setoid_rewrite_left rinstInst'_Tm;
+                  try setoid_rewrite_left rinstInst'_PTm_pointwise;
+                  try setoid_rewrite_left rinstInst'_PTm.
 
 End Core.
 
@@ -945,12 +1684,8 @@ Arguments Bool {n_Tm}.
 
 Arguments BVal {n_Tm}.
 
-Arguments Bot {n_Tm}.
-
 Arguments Univ {n_Tm}.
 
-Arguments Const {n_Tm}.
-
 Arguments TBind {n_Tm}.
 
 Arguments Proj {n_Tm}.
@@ -961,10 +1696,30 @@ Arguments App {n_Tm}.
 
 Arguments Abs {n_Tm}.
 
+Arguments VarPTm {n_PTm}.
+
+Arguments PBot {n_PTm}.
+
+Arguments PUniv {n_PTm}.
+
+Arguments PConst {n_PTm}.
+
+Arguments PProj {n_PTm}.
+
+Arguments PPair {n_PTm}.
+
+Arguments PApp {n_PTm}.
+
+Arguments PAbs {n_PTm}.
+
 #[global]Hint Opaque subst_Tm: rewrite.
 
 #[global]Hint Opaque ren_Tm: rewrite.
 
+#[global]Hint Opaque subst_PTm: rewrite.
+
+#[global]Hint Opaque ren_PTm: rewrite.
+
 End Extra.
 
 Module interface.
diff --git a/theories/diagram.txt b/theories/diagram.txt
index 2cf16e8..ab47f3d 100644
--- a/theories/diagram.txt
+++ b/theories/diagram.txt
@@ -18,3 +18,25 @@ a0 >> a1
 |     |
 v     v
 b0 >> b1
+
+
+prov x (x, x)
+
+prov x b
+
+
+a => b
+
+prov x a
+
+prov y b
+
+prov x c
+prov y c
+
+extract c = x
+extract c = y
+
+prov x b
+
+pr
diff --git a/theories/fp_red.v b/theories/fp_red.v
index e9b5591..3482f45 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -22,86 +22,82 @@ Ltac spec_refl := ltac2:(spec_refl ()).
 
 (* Trying my best to not write C style module_funcname *)
 Module Par.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (***************** Beta ***********************)
   | AppAbs a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App (Abs a0) b0)  (subst_Tm (scons b1 VarTm) a1)
+    R (PApp (PAbs a0) b0)  (subst_PTm (scons b1 VarPTm) a1)
   | AppPair a0 a1 b0 b1 c0 c1:
     R a0 a1 ->
     R b0 b1 ->
     R c0 c1 ->
-    R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1))
+    R (PApp (PPair a0 b0) c0) (PPair (PApp a1 c1) (PApp b1 c1))
   | ProjAbs p a0 a1 :
     R a0 a1 ->
-    R (Proj p (Abs a0)) (Abs (Proj p a1))
+    R (PProj p (PAbs a0)) (PAbs (PProj p a1))
   | ProjPair p a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
+    R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1)
 
   (****************** Eta ***********************)
   | AppEta a0 a1 :
     R a0 a1 ->
-    R a0 (Abs (App (ren_Tm shift a1) (VarTm var_zero)))
+    R a0 (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero)))
   | PairEta a0 a1 :
     R a0 a1 ->
-    R a0 (Pair (Proj PL a1) (Proj PR a1))
+    R a0 (PPair (PProj PL a1) (PProj PR a1))
 
   (*************** Congruence ********************)
-  | Var i : R (VarTm i) (VarTm i)
+  | Var i : R (VarPTm i) (VarPTm i)
   | AbsCong a0 a1 :
     R a0 a1 ->
-    R (Abs a0) (Abs a1)
+    R (PAbs a0) (PAbs a1)
   | AppCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App a0 b0) (App a1 b1)
+    R (PApp a0 b0) (PApp a1 b1)
   | PairCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Pair a0 b0) (Pair a1 b1)
+    R (PPair a0 b0) (PPair a1 b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong p A0 A1 B0 B1:
-    R A0 A1 ->
-    R B0 B1 ->
-    R (TBind p A0 B0) (TBind p A1 B1)
+    R (PProj p a0) (PProj p a1)
   | ConstCong k :
-    R (Const k) (Const k)
-  | UnivCong i :
-    R (Univ i) (Univ i)
-  | BotCong :
-    R Bot Bot.
+    R (PConst k) (PConst k)
+  | Univ i :
+    R (PUniv i) (PUniv i)
+  | Bot :
+    R PBot PBot.
 
-  Lemma refl n (a : Tm n) : R a a.
+  Lemma refl n (a : PTm n) : R a a.
     elim : n /a; hauto ctrs:R.
   Qed.
 
-  Lemma AppAbs' n a0 a1 (b0 b1 t : Tm n) :
-    t = subst_Tm (scons b1 VarTm) a1 ->
+  Lemma AppAbs' n a0 a1 (b0 b1 t : PTm n) :
+    t = subst_PTm (scons b1 VarPTm) a1 ->
     R a0 a1 ->
     R b0 b1 ->
-    R (App (Abs a0) b0) t.
+    R (PApp (PAbs a0) b0) t.
   Proof. move => ->. apply AppAbs. Qed.
 
-  Lemma ProjPair' n p (a0 a1 b0 b1 : Tm n) t :
+  Lemma ProjPair' n p (a0 a1 b0 b1 : PTm n) t :
     t = (if p is PL then a1 else b1) ->
     R a0 a1 ->
     R b0 b1 ->
-    R (Proj p (Pair a0 b0)) t.
+    R (PProj p (PPair a0 b0)) t.
   Proof.  move => > ->. apply ProjPair. Qed.
 
-  Lemma AppEta' n (a0 a1 b : Tm n) :
-    b = (Abs (App (ren_Tm shift a1) (VarTm var_zero))) ->
+  Lemma AppEta' n (a0 a1 b : PTm n) :
+    b = (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) ->
     R a0 a1 ->
     R a0 b.
   Proof. move => ->; apply AppEta. Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
@@ -113,13 +109,13 @@ Module Par.
   Qed.
 
 
-  Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
+  Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    R a b -> R (subst_Tm ρ0 a) (subst_Tm ρ1 b).
+    R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
   Proof.
     move => + h. move : m ρ0 ρ1. elim : n a b/h.
     - move => n a0 a1 b0 b1 ha iha hb ihb m ρ0 ρ1 hρ /=.
-      eapply AppAbs' with (a1 := subst_Tm (up_Tm_Tm ρ1) a1); eauto.
+      eapply AppAbs' with (a1 := subst_PTm (up_PTm_PTm ρ1) a1); eauto.
       by asimpl.
       hauto l:on use:renaming inv:option.
     - hauto lq:on rew:off ctrs:R.
@@ -134,19 +130,18 @@ Module Par.
     - qauto l:on ctrs:R.
     - qauto l:on ctrs:R.
     - hauto l:on inv:option ctrs:R use:renaming.
-    - sfirstorder.
-    - sfirstorder.
-    - sfirstorder.
+    - qauto l:on ctrs:R.
+    - qauto l:on ctrs:R.
   Qed.
 
-  Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
-    R a b -> R (subst_Tm ρ a) (subst_Tm ρ b).
+  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof. hauto l:on use:morphing, refl. Qed.
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ξ : fin n -> fin m) :
-    R (ren_Tm ξ a) b -> exists b0, R a b0 /\ ren_Tm ξ b0 = b.
+  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
+    R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
   Proof.
-    move E : (ren_Tm ξ a) => u h.
+    move E : (ren_PTm ξ a) => u h.
     move : n ξ a E. elim : m u b/h.
     - move => n a0 a1 b0 b1 ha iha hb ihb m ξ []//=.
       move => c c0 [+ ?]. subst.
@@ -190,7 +185,7 @@ Module Par.
       spec_refl.
       move  :iha => [b0 [? ?]]. subst.
       eexists. split. by apply AbsCong; eauto.
-      by asimpl.
+      done.
     - move => n a0 a1 b0 b1 ha iha hb ihb m ξ []//= t t0 [*]. subst.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
@@ -201,43 +196,37 @@ Module Par.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       move : ihb => [c0 [? ?]]. subst.
-      eexists. split. by apply PairCong; eauto.
-      by asimpl.
+      eexists. split=>/=. by apply PairCong; eauto.
+      done.
     - move => n p a0 a1 ha iha m ξ []//= p0 t [*]. subst.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       eexists. split. by apply ProjCong; eauto.
-      by asimpl.
-    - move => n p A0 A1 B0 B1 ha iha hB ihB m ξ []//= ? t t0 [*]. subst.
-      spec_refl.
-      move : iha => [b0 [? ?]].
-      move : ihB => [c0 [? ?]]. subst.
-      eexists. split. by apply BindCong; eauto.
-      by asimpl.
-    - move => n k m ξ []//=. hauto l:on.
-    - move => n i n0 ξ []//=. hauto l:on.
-    - hauto q:on inv:Tm ctrs:R.
+      done.
+    - hauto q:on inv:PTm ctrs:R.
+    - hauto q:on inv:PTm ctrs:R.
+    - hauto q:on inv:PTm ctrs:R.
   Qed.
 
 End Par.
 
 Module Pars.
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    rtc Par.R a b -> rtc Par.R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    rtc Par.R a b -> rtc Par.R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     induction 1; hauto lq:on ctrs:rtc use:Par.renaming.
   Qed.
 
-  Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
     rtc Par.R a b ->
-    rtc Par.R (subst_Tm ρ a) (subst_Tm ρ b).
+    rtc Par.R (subst_PTm ρ a) (subst_PTm ρ b).
     induction 1; hauto l:on ctrs:rtc use:Par.substing.
   Qed.
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ξ : fin n -> fin m) :
-    rtc Par.R (ren_Tm ξ a) b -> exists b0, rtc Par.R a b0 /\ ren_Tm ξ b0 = b.
+  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
+    rtc Par.R (ren_PTm ξ a) b -> exists b0, rtc Par.R a b0 /\ ren_PTm ξ b0 = b.
   Proof.
-    move E  :(ren_Tm ξ a) => u h.
+    move E  :(ren_PTm ξ a) => u h.
     move : a E.
     elim : u b /h.
     - sfirstorder.
@@ -254,109 +243,106 @@ Module Pars.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma ProjCong n p (a0 a1 : Tm n) :
+  Lemma ProjCong n p (a0 a1 : PTm n) :
     rtc Par.R a0 a1 ->
-    rtc Par.R (Proj p a0) (Proj p a1).
+    rtc Par.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     rtc Par.R a0 a1 ->
     rtc Par.R b0 b1 ->
-    rtc Par.R (Pair a0 b0) (Pair a1 b1).
+    rtc Par.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     rtc Par.R a0 a1 ->
     rtc Par.R b0 b1 ->
-    rtc Par.R (App a0 b0) (App a1 b1).
+    rtc Par.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma AbsCong n (a b : Tm (S n)) :
+  Lemma AbsCong n (a b : PTm (S n)) :
     rtc Par.R a b ->
-    rtc Par.R (Abs a) (Abs b).
+    rtc Par.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
 End Pars.
 
-Definition var_or_const {n} (a : Tm n) :=
+Definition var_or_const {n} (a : PTm n) :=
   match a with
-  | VarTm _ => true
-  | Bot => true
+  | VarPTm _ => true
+  | PBot => true
   | _ => false
   end.
 
+
 (***************** Beta rules only ***********************)
 Module RPar.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (***************** Beta ***********************)
   | AppAbs a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App (Abs a0) b0)  (subst_Tm (scons b1 VarTm) a1)
+    R (PApp (PAbs a0) b0)  (subst_PTm (scons b1 VarPTm) a1)
   | AppPair a0 a1 b0 b1 c0 c1:
     R a0 a1 ->
     R b0 b1 ->
     R c0 c1 ->
-    R (App (Pair a0 b0) c0) (Pair (App a1 c1) (App b1 c1))
+    R (PApp (PPair a0 b0) c0) (PPair (PApp a1 c1) (PApp b1 c1))
   | ProjAbs p a0 a1 :
     R a0 a1 ->
-    R (Proj p (Abs a0)) (Abs (Proj p a1))
+    R (PProj p (PAbs a0)) (PAbs (PProj p a1))
   | ProjPair p a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
+    R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1)
 
 
   (*************** Congruence ********************)
-  | Var i : R (VarTm i) (VarTm i)
+  | Var i : R (VarPTm i) (VarPTm i)
   | AbsCong a0 a1 :
     R a0 a1 ->
-    R (Abs a0) (Abs a1)
+    R (PAbs a0) (PAbs a1)
   | AppCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App a0 b0) (App a1 b1)
+    R (PApp a0 b0) (PApp a1 b1)
   | PairCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Pair a0 b0) (Pair a1 b1)
+    R (PPair a0 b0) (PPair a1 b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong p A0 A1 B0 B1:
-    R A0 A1 ->
-    R B0 B1 ->
-    R (TBind p A0 B0) (TBind p A1 B1)
+    R (PProj p a0) (PProj p a1)
   | ConstCong k :
-    R (Const k) (Const k)
-  | UnivCong i :
-    R (Univ i) (Univ i)
-  | BotCong :
-    R Bot Bot.
+    R (PConst k) (PConst k)
+  | Univ i :
+    R (PUniv i) (PUniv i)
+  | Bot :
+    R PBot PBot.
 
-  Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
+  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
-  Lemma refl n (a : Tm n) : R a a.
+  Lemma refl n (a : PTm n) : R a a.
   Proof.
     induction a; hauto lq:on ctrs:R.
   Qed.
 
-  Lemma AppAbs' n a0 a1 (b0 b1 t : Tm n) :
-    t = subst_Tm (scons b1 VarTm) a1 ->
+  Lemma AppAbs' n a0 a1 (b0 b1 t : PTm n) :
+    t = subst_PTm (scons b1 VarPTm) a1 ->
     R a0 a1 ->
     R b0 b1 ->
-    R (App (Abs a0) b0) t.
+    R (PApp (PAbs a0) b0) t.
   Proof. move => ->. apply AppAbs. Qed.
 
-  Lemma ProjPair' n p (a0 a1 b0 b1 : Tm n) t :
+  Lemma ProjPair' n p (a0 a1 b0 b1 : PTm n) t :
     t = (if p is PL then a1 else b1) ->
     R a0 a1 ->
     R b0 b1 ->
-    R (Proj p (Pair a0 b0)) t.
+    R (PProj p (PPair a0 b0)) t.
   Proof.  move => > ->. apply ProjPair. Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
@@ -364,25 +350,25 @@ Module RPar.
     all : qauto ctrs:R use:ProjPair'.
   Qed.
 
-  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> Tm m) (ξ : fin m -> fin p) :
+  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R ((funcomp (ren_Tm ξ) ρ0) i) ((funcomp (ren_Tm ξ) ρ1) i)).
+    (forall i, R ((funcomp (ren_PTm ξ) ρ0) i) ((funcomp (ren_PTm ξ) ρ1) i)).
   Proof. eauto using renaming. Qed.
 
-  Lemma morphing_ext n m (ρ0 ρ1 : fin n -> Tm m) a b  :
+  Lemma morphing_ext n m (ρ0 ρ1 : fin n -> PTm m) a b  :
     R a b ->
     (forall i, R (ρ0 i) (ρ1 i)) ->
     (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)).
   Proof. hauto q:on inv:option. Qed.
 
-  Lemma morphing_up n m (ρ0 ρ1 : fin n -> Tm m) :
+  Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R (up_Tm_Tm ρ0 i) (up_Tm_Tm ρ1 i)).
-  Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_Tm_Tm. Qed.
+    (forall i, R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)).
+  Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_PTm_PTm. Qed.
 
-  Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
+  Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    R a b -> R (subst_Tm ρ0 a) (subst_Tm ρ1 b).
+    R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
   Proof.
     move => + h. move : m ρ0 ρ1.
     elim : n a b /h.
@@ -400,33 +386,32 @@ Module RPar.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
-    - hauto lq:on ctrs:R.
   Qed.
 
-  Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
     R a b ->
-    R (subst_Tm ρ a) (subst_Tm ρ b).
+    R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof. hauto l:on use:morphing, refl. Qed.
 
-  Lemma cong n (a b : Tm (S n)) c d :
+  Lemma cong n (a b : PTm (S n)) c d :
     R a b ->
     R c d ->
-    R (subst_Tm (scons c VarTm) a) (subst_Tm (scons d VarTm) b).
+    R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons d VarPTm) b).
   Proof.
     move => h0 h1. apply morphing => //=.
     qauto l:on ctrs:R inv:option.
   Qed.
 
-  Lemma var_or_const_imp {n} (a b : Tm n) :
+  Lemma var_or_const_imp {n} (a b : PTm n) :
     var_or_const a ->
     a = b -> ~~ var_or_const b -> False.
   Proof.
-    hauto lq:on inv:Tm.
+    hauto lq:on inv:PTm.
   Qed.
 
-  Lemma var_or_const_up n m (ρ : fin n -> Tm m) :
+  Lemma var_or_const_up n m (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
-    (forall i, var_or_const (up_Tm_Tm ρ i)).
+    (forall i, var_or_const (up_PTm_PTm ρ i)).
   Proof.
     move => h /= [i|].
     - asimpl.
@@ -439,11 +424,11 @@ Module RPar.
 
   Local Ltac antiimp := qauto l:on use:var_or_const_imp.
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
+  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
-    R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b.
+    R (subst_PTm ρ a) b -> exists b0, R a b0 /\ subst_PTm ρ b0 = b.
   Proof.
-    move E : (subst_Tm ρ a) => u hρ h.
+    move E : (subst_PTm ρ a) => u hρ h.
     move : n ρ hρ a E. elim : m u b/h.
     - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
                first by antiimp.
@@ -460,7 +445,8 @@ Module RPar.
       eexists. split.
       apply AppAbs; eauto.
       by asimpl.
-    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ hρ []//=;
+    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ρ hρ.
+      move => []//=;
                first by antiimp.
       move => []//=; first by antiimp.
       move => t t0 t1 [*]. subst.
@@ -527,87 +513,72 @@ Module RPar.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       eexists. split. apply ProjCong; eauto. reflexivity.
-    - move => n p A0 A1 B0 B1 ha iha hB ihB m ρ hρ []//=;
-               first by antiimp.
-      move => ? t t0 [*]. subst.
-      have {}/iha := (hρ) => iha.
-      have /var_or_const_up {}/ihB := (hρ) => ihB.
-      spec_refl.
-      move : iha => [b0 [? ?]].
-      move : ihB => [c0 [? ?]]. subst.
-      eexists. split. by apply BindCong; eauto.
-      by asimpl.
-    - hauto q:on ctrs:R inv:Tm.
-    - move => n i n0 ρ hρ []//=; first by antiimp.
-      hauto l:on.
-    - move => n m ρ hρ []//=; hauto lq:on ctrs:R.
+    - hauto q:on ctrs:R inv:PTm.
+    - hauto q:on ctrs:R inv:PTm.
+    - hauto q:on ctrs:R inv:PTm.
   Qed.
 End RPar.
 
 (***************** Beta rules only ***********************)
 Module RPar'.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (***************** Beta ***********************)
   | AppAbs a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App (Abs a0) b0)  (subst_Tm (scons b1 VarTm) a1)
+    R (PApp (PAbs a0) b0)  (subst_PTm (scons b1 VarPTm) a1)
   | ProjPair p a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Proj p (Pair a0 b0)) (if p is PL then a1 else b1)
+    R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1)
 
 
   (*************** Congruence ********************)
-  | Var i : R (VarTm i) (VarTm i)
+  | Var i : R (VarPTm i) (VarPTm i)
   | AbsCong a0 a1 :
     R a0 a1 ->
-    R (Abs a0) (Abs a1)
+    R (PAbs a0) (PAbs a1)
   | AppCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App a0 b0) (App a1 b1)
+    R (PApp a0 b0) (PApp a1 b1)
   | PairCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Pair a0 b0) (Pair a1 b1)
+    R (PPair a0 b0) (PPair a1 b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong p A0 A1 B0 B1:
-    R A0 A1 ->
-    R B0 B1 ->
-    R (TBind p A0 B0) (TBind p A1 B1)
+    R (PProj p a0) (PProj p a1)
   | ConstCong k :
-    R (Const k) (Const k)
+    R (PConst k) (PConst k)
   | UnivCong i :
-    R (Univ i) (Univ i)
+    R (PUniv i) (PUniv i)
   | BotCong :
-    R Bot Bot.
+    R PBot PBot.
 
-  Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
+  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
-  Lemma refl n (a : Tm n) : R a a.
+  Lemma refl n (a : PTm n) : R a a.
   Proof.
     induction a; hauto lq:on ctrs:R.
   Qed.
 
-  Lemma AppAbs' n a0 a1 (b0 b1 t : Tm n) :
-    t = subst_Tm (scons b1 VarTm) a1 ->
+  Lemma AppAbs' n a0 a1 (b0 b1 t : PTm n) :
+    t = subst_PTm (scons b1 VarPTm) a1 ->
     R a0 a1 ->
     R b0 b1 ->
-    R (App (Abs a0) b0) t.
+    R (PApp (PAbs a0) b0) t.
   Proof. move => ->. apply AppAbs. Qed.
 
-  Lemma ProjPair' n p (a0 a1 b0 b1 : Tm n) t :
+  Lemma ProjPair' n p (a0 a1 b0 b1 : PTm n) t :
     t = (if p is PL then a1 else b1) ->
     R a0 a1 ->
     R b0 b1 ->
-    R (Proj p (Pair a0 b0)) t.
+    R (PProj p (PPair a0 b0)) t.
   Proof.  move => > ->. apply ProjPair. Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
@@ -615,25 +586,25 @@ Module RPar'.
     all : qauto ctrs:R use:ProjPair'.
   Qed.
 
-  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> Tm m) (ξ : fin m -> fin p) :
+  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R ((funcomp (ren_Tm ξ) ρ0) i) ((funcomp (ren_Tm ξ) ρ1) i)).
+    (forall i, R ((funcomp (ren_PTm ξ) ρ0) i) ((funcomp (ren_PTm ξ) ρ1) i)).
   Proof. eauto using renaming. Qed.
 
-  Lemma morphing_ext n m (ρ0 ρ1 : fin n -> Tm m) a b  :
+  Lemma morphing_ext n m (ρ0 ρ1 : fin n -> PTm m) a b  :
     R a b ->
     (forall i, R (ρ0 i) (ρ1 i)) ->
     (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)).
   Proof. hauto q:on inv:option. Qed.
 
-  Lemma morphing_up n m (ρ0 ρ1 : fin n -> Tm m) :
+  Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R (up_Tm_Tm ρ0 i) (up_Tm_Tm ρ1 i)).
-  Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_Tm_Tm. Qed.
+    (forall i, R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)).
+  Proof. hauto l:on ctrs:R use:morphing_ext, morphing_ren unfold:up_PTm_PTm. Qed.
 
-  Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
+  Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    R a b -> R (subst_Tm ρ0 a) (subst_Tm ρ1 b).
+    R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
   Proof.
     move => + h. move : m ρ0 ρ1.
     elim : n a b /h.
@@ -646,36 +617,35 @@ Module RPar'.
     - hauto lq:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
-    - hauto lq:on ctrs:R use:morphing_up.
-    - hauto lq:on ctrs:R.
+    - hauto l:on ctrs:R use:morphing_up.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
   Qed.
 
-  Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
     R a b ->
-    R (subst_Tm ρ a) (subst_Tm ρ b).
+    R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof. hauto l:on use:morphing, refl. Qed.
 
-  Lemma cong n (a b : Tm (S n)) c d :
+  Lemma cong n (a b : PTm (S n)) c d :
     R a b ->
     R c d ->
-    R (subst_Tm (scons c VarTm) a) (subst_Tm (scons d VarTm) b).
+    R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons d VarPTm) b).
   Proof.
     move => h0 h1. apply morphing => //=.
     qauto l:on ctrs:R inv:option.
   Qed.
 
-  Lemma var_or_const_imp {n} (a b : Tm n) :
+  Lemma var_or_const_imp {n} (a b : PTm n) :
     var_or_const a ->
     a = b -> ~~ var_or_const b -> False.
   Proof.
-    hauto lq:on inv:Tm.
+    hauto lq:on inv:PTm.
   Qed.
 
-  Lemma var_or_const_up n m (ρ : fin n -> Tm m) :
+  Lemma var_or_const_up n m (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
-    (forall i, var_or_const (up_Tm_Tm ρ i)).
+    (forall i, var_or_const (up_PTm_PTm ρ i)).
   Proof.
     move => h /= [i|].
     - asimpl.
@@ -688,11 +658,11 @@ Module RPar'.
 
   Local Ltac antiimp := qauto l:on use:var_or_const_imp.
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
+  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
-    R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b.
+    R (subst_PTm ρ a) b -> exists b0, R a b0 /\ subst_PTm ρ b0 = b.
   Proof.
-    move E : (subst_Tm ρ a) => u hρ h.
+    move E : (subst_PTm ρ a) => u hρ h.
     move : n ρ hρ a E. elim : m u b/h.
     - move => n a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
                first by antiimp.
@@ -756,66 +726,50 @@ Module RPar'.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       eexists. split. apply ProjCong; eauto. reflexivity.
-    - move => n p A0 A1 B0 B1 ha iha hB ihB m ρ hρ []//=;
-               first by antiimp.
-      move => ? t t0 [*]. subst.
-      have {}/iha := (hρ) => iha.
-      have /var_or_const_up {}/ihB := (hρ) => ihB.
-      spec_refl.
-      move : iha => [b0 [? ?]].
-      move : ihB => [c0 [? ?]]. subst.
-      eexists. split. by apply BindCong; eauto.
-      by asimpl.
-    - hauto q:on ctrs:R inv:Tm.
+    - hauto q:on ctrs:R inv:PTm.
     - move => n i n0 ρ hρ []//=; first by antiimp.
       hauto l:on.
-    - hauto q:on inv:Tm ctrs:R.
+    - hauto q:on inv:PTm ctrs:R.
   Qed.
 End RPar'.
 
 Module ERed.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (****************** Eta ***********************)
   | AppEta a :
-    R a (Abs (App (ren_Tm shift a) (VarTm var_zero)))
+    R a (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero)))
   | PairEta a :
-    R a (Pair (Proj PL a) (Proj PR a))
+    R a (PPair (PProj PL a) (PProj PR a))
 
   (*************** Congruence ********************)
   | AbsCong a0 a1 :
     R a0 a1 ->
-    R (Abs a0) (Abs a1)
+    R (PAbs a0) (PAbs a1)
   | AppCong0 a0 a1 b :
     R a0 a1 ->
-    R (App a0 b) (App a1 b)
+    R (PApp a0 b) (PApp a1 b)
   | AppCong1 a b0 b1 :
     R b0 b1 ->
-    R (App a b0) (App a b1)
+    R (PApp a b0) (PApp a b1)
   | PairCong0 a0 a1 b :
     R a0 a1 ->
-    R (Pair a0 b) (Pair a1 b)
+    R (PPair a0 b) (PPair a1 b)
   | PairCong1 a b0 b1 :
     R b0 b1 ->
-    R (Pair a b0) (Pair a b1)
+    R (PPair a b0) (PPair a b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong0 p A0 A1 B:
-    R A0 A1 ->
-    R (TBind p A0 B) (TBind p A1 B)
-  | BindCong1 p A B0 B1:
-    R B0 B1 ->
-    R (TBind p A B0) (TBind p A B1).
+    R (PProj p a0) (PProj p a1).
 
-  Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
+  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
-  Lemma AppEta' n a (u : Tm n) :
-    u = (Abs (App (ren_Tm shift a) (VarTm var_zero))) ->
+  Lemma AppEta' n a (u : PTm n) :
+    u = (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) ->
     R a u.
   Proof. move => ->. apply AppEta. Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
@@ -825,9 +779,9 @@ Module ERed.
     all : qauto ctrs:R.
   Qed.
 
-  Lemma substing n m (a : Tm n) b (ρ : fin n -> Tm m) :
+  Lemma substing n m (a : PTm n) b (ρ : fin n -> PTm m) :
     R a b ->
-    R (subst_Tm ρ a) (subst_Tm ρ b).
+    R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     move => h. move : m ρ. elim : n a b / h => n.
     move => a m ρ /=.
@@ -846,79 +800,69 @@ Module EReds.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AbsCong n (a b : Tm (S n)) :
+  Lemma AbsCong n (a b : PTm (S n)) :
     rtc ERed.R a b ->
-    rtc ERed.R (Abs a) (Abs b).
+    rtc ERed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
-    rtc ERed.R (App a0 b0) (App a1 b1).
+    rtc ERed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1 :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
-    rtc ERed.R (TBind p a0 b0) (TBind p a1 b1).
+    rtc ERed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma ProjCong n p (a0 a1  : PTm n) :
     rtc ERed.R a0 a1 ->
-    rtc ERed.R b0 b1 ->
-    rtc ERed.R (Pair a0 b0) (Pair a1 b1).
-  Proof. solve_s. Qed.
-
-  Lemma ProjCong n p (a0 a1  : Tm n) :
-    rtc ERed.R a0 a1 ->
-    rtc ERed.R (Proj p a0) (Proj p a1).
+    rtc ERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 End EReds.
 
 Module EPar.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (****************** Eta ***********************)
   | AppEta a0 a1 :
     R a0 a1 ->
-    R a0 (Abs (App (ren_Tm shift a1) (VarTm var_zero)))
+    R a0 (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero)))
   | PairEta a0 a1 :
     R a0 a1 ->
-    R a0 (Pair (Proj PL a1) (Proj PR a1))
+    R a0 (PPair (PProj PL a1) (PProj PR a1))
 
   (*************** Congruence ********************)
-  | Var i : R (VarTm i) (VarTm i)
+  | Var i : R (VarPTm i) (VarPTm i)
   | AbsCong a0 a1 :
     R a0 a1 ->
-    R (Abs a0) (Abs a1)
+    R (PAbs a0) (PAbs a1)
   | AppCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (App a0 b0) (App a1 b1)
+    R (PApp a0 b0) (PApp a1 b1)
   | PairCong a0 a1 b0 b1 :
     R a0 a1 ->
     R b0 b1 ->
-    R (Pair a0 b0) (Pair a1 b1)
+    R (PPair a0 b0) (PPair a1 b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (Proj p a0) (Proj p a1)
-  | BindCong p A0 A1 B0 B1:
-    R A0 A1 ->
-    R B0 B1 ->
-    R (TBind p A0 B0) (TBind p A1 B1)
+    R (PProj p a0) (PProj p a1)
   | ConstCong k :
-    R (Const k) (Const k)
+    R (PConst k) (PConst k)
   | UnivCong i :
-    R (Univ i) (Univ i)
+    R (PUniv i) (PUniv i)
   | BotCong :
-    R Bot Bot.
+    R PBot PBot.
 
-  Lemma refl n (a : Tm n) : EPar.R a a.
+  Lemma refl n (a : PTm n) : EPar.R a a.
   Proof.
     induction a; hauto lq:on ctrs:EPar.R.
   Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
@@ -930,18 +874,18 @@ Module EPar.
     all : qauto ctrs:R.
   Qed.
 
-  Derive Dependent Inversion inv with (forall n (a b : Tm n), R a b) Sort Prop.
+  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
-  Lemma AppEta' n (a0 a1 b : Tm n) :
-    b = (Abs (App (ren_Tm shift a1) (VarTm var_zero))) ->
+  Lemma AppEta' n (a0 a1 b : PTm n) :
+    b = (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) ->
     R a0 a1 ->
     R a0 b.
   Proof. move => ->; apply AppEta. Qed.
 
-  Lemma morphing n m (a b : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
+  Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
     R a b ->
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    R (subst_Tm ρ0 a) (subst_Tm ρ1 b).
+    R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
   Proof.
     move => h. move : m ρ0 ρ1. elim : n a b / h => n.
     - move => a0 a1 ha iha m ρ0 ρ1 hρ /=.
@@ -955,13 +899,12 @@ Module EPar.
     - hauto l:on ctrs:R use:renaming inv:option.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
-    - hauto lq:on ctrs:R.
   Qed.
 
-  Lemma substing n a0 a1 (b0 b1 : Tm n) :
+  Lemma substing n a0 a1 (b0 b1 : PTm n) :
     R a0 a1 ->
     R b0 b1 ->
-    R (subst_Tm (scons b0 VarTm) a0) (subst_Tm (scons b1 VarTm) a1).
+    R (subst_PTm (scons b0 VarPTm) a0) (subst_PTm (scons b1 VarPTm) a1).
   Proof.
     move => h0 h1. apply morphing => //.
     hauto lq:on ctrs:R inv:option.
@@ -971,14 +914,14 @@ End EPar.
 
 
 Module OExp.
-  Inductive R {n} : Tm n -> Tm n ->  Prop :=
+  Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (****************** Eta ***********************)
   | AppEta a :
-    R a (Abs (App (ren_Tm shift a) (VarTm var_zero)))
+    R a (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero)))
   | PairEta a :
-    R a (Pair (Proj PL a) (Proj PR a)).
+    R a (PPair (PProj PL a) (PProj PR a)).
 
-  Lemma merge n (t a b : Tm n) :
+  Lemma merge n (t a b : PTm n) :
     rtc R a b ->
     EPar.R t a ->
     EPar.R t b.
@@ -988,7 +931,7 @@ Module OExp.
     - hauto q:on ctrs:EPar.R inv:R.
   Qed.
 
-  Lemma commutativity n (a b c : Tm n) :
+  Lemma commutativity n (a b c : PTm n) :
     EPar.R a b -> R a c -> exists d, R b d /\ EPar.R c d.
   Proof.
     move => h.
@@ -997,7 +940,7 @@ Module OExp.
     - hauto lq:on ctrs:EPar.R, R.
   Qed.
 
-  Lemma commutativity0 n (a b c : Tm n) :
+  Lemma commutativity0 n (a b c : PTm n) :
     EPar.R a b -> rtc R a c -> exists d, rtc R b d /\ EPar.R c d.
   Proof.
     move => + h. move : b.
@@ -1022,72 +965,66 @@ Module RPars.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AbsCong n (a b : Tm (S n)) :
+  Lemma AbsCong n (a b : PTm (S n)) :
     rtc RPar.R a b ->
-    rtc RPar.R (Abs a) (Abs b).
+    rtc RPar.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     rtc RPar.R a0 a1 ->
     rtc RPar.R b0 b1 ->
-    rtc RPar.R (App a0 b0) (App a1 b1).
+    rtc RPar.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1 :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     rtc RPar.R a0 a1 ->
     rtc RPar.R b0 b1 ->
-    rtc RPar.R (TBind p a0 b0) (TBind p a1 b1).
+    rtc RPar.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma ProjCong n p (a0 a1  : PTm n) :
     rtc RPar.R a0 a1 ->
-    rtc RPar.R b0 b1 ->
-    rtc RPar.R (Pair a0 b0) (Pair a1 b1).
+    rtc RPar.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : Tm n) :
+  Lemma renaming n (a0 a1 : PTm n) m (ξ : fin n -> fin m) :
     rtc RPar.R a0 a1 ->
-    rtc RPar.R (Proj p a0) (Proj p a1).
-  Proof. solve_s. Qed.
-
-  Lemma renaming n (a0 a1 : Tm n) m (ξ : fin n -> fin m) :
-    rtc RPar.R a0 a1 ->
-    rtc RPar.R (ren_Tm ξ a0) (ren_Tm ξ a1).
+    rtc RPar.R (ren_PTm ξ a0) (ren_PTm ξ a1).
   Proof.
     induction 1.
     - apply rtc_refl.
     - eauto using RPar.renaming, rtc_l.
   Qed.
 
-  Lemma weakening n (a0 a1 : Tm n) :
+  Lemma weakening n (a0 a1 : PTm n) :
     rtc RPar.R a0 a1 ->
-    rtc RPar.R (ren_Tm shift a0) (ren_Tm shift a1).
+    rtc RPar.R (ren_PTm shift a0) (ren_PTm shift a1).
   Proof. apply renaming. Qed.
 
-  Lemma Abs_inv n (a : Tm (S n)) b :
-    rtc RPar.R (Abs a) b -> exists a', b = Abs a' /\ rtc RPar.R a a'.
+  Lemma Abs_inv n (a : PTm (S n)) b :
+    rtc RPar.R (PAbs a) b -> exists a', b = PAbs a' /\ rtc RPar.R a a'.
   Proof.
-    move E : (Abs a) => b0 h. move : a E.
+    move E : (PAbs a) => b0 h. move : a E.
     elim : b0 b / h.
     - hauto lq:on ctrs:rtc.
     - hauto lq:on ctrs:rtc inv:RPar.R, rtc.
   Qed.
 
-  Lemma morphing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+  Lemma morphing n m (a b : PTm n) (ρ : fin n -> PTm m) :
     rtc RPar.R a b ->
-    rtc RPar.R (subst_Tm ρ a) (subst_Tm ρ b).
+    rtc RPar.R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof. induction 1; qauto l:on ctrs:rtc use:RPar.substing. Qed.
 
-  Lemma substing n (a b : Tm (S n)) c :
+  Lemma substing n (a b : PTm (S n)) c :
     rtc RPar.R a b ->
-    rtc RPar.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b).
+    rtc RPar.R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons c VarPTm) b).
   Proof. hauto lq:on use:morphing inv:option. Qed.
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
+  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
-    rtc RPar.R (subst_Tm ρ a) b -> exists b0, rtc RPar.R a b0 /\ subst_Tm ρ b0 = b.
+    rtc RPar.R (subst_PTm ρ a) b -> exists b0, rtc RPar.R a b0 /\ subst_PTm ρ b0 = b.
   Proof.
-    move E  :(subst_Tm ρ a) => u hρ h.
+    move E  :(subst_PTm ρ a) => u hρ h.
     move : a E.
     elim : u b /h.
     - sfirstorder.
@@ -1109,72 +1046,66 @@ Module RPars'.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AbsCong n (a b : Tm (S n)) :
+  Lemma AbsCong n (a b : PTm (S n)) :
     rtc RPar'.R a b ->
-    rtc RPar'.R (Abs a) (Abs b).
+    rtc RPar'.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     rtc RPar'.R a0 a1 ->
     rtc RPar'.R b0 b1 ->
-    rtc RPar'.R (App a0 b0) (App a1 b1).
+    rtc RPar'.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1 :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     rtc RPar'.R a0 a1 ->
     rtc RPar'.R b0 b1 ->
-    rtc RPar'.R (TBind p a0 b0) (TBind p a1 b1).
+    rtc RPar'.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma ProjCong n p (a0 a1  : PTm n) :
     rtc RPar'.R a0 a1 ->
-    rtc RPar'.R b0 b1 ->
-    rtc RPar'.R (Pair a0 b0) (Pair a1 b1).
+    rtc RPar'.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : Tm n) :
+  Lemma renaming n (a0 a1 : PTm n) m (ξ : fin n -> fin m) :
     rtc RPar'.R a0 a1 ->
-    rtc RPar'.R (Proj p a0) (Proj p a1).
-  Proof. solve_s. Qed.
-
-  Lemma renaming n (a0 a1 : Tm n) m (ξ : fin n -> fin m) :
-    rtc RPar'.R a0 a1 ->
-    rtc RPar'.R (ren_Tm ξ a0) (ren_Tm ξ a1).
+    rtc RPar'.R (ren_PTm ξ a0) (ren_PTm ξ a1).
   Proof.
     induction 1.
     - apply rtc_refl.
     - eauto using RPar'.renaming, rtc_l.
   Qed.
 
-  Lemma weakening n (a0 a1 : Tm n) :
+  Lemma weakening n (a0 a1 : PTm n) :
     rtc RPar'.R a0 a1 ->
-    rtc RPar'.R (ren_Tm shift a0) (ren_Tm shift a1).
+    rtc RPar'.R (ren_PTm shift a0) (ren_PTm shift a1).
   Proof. apply renaming. Qed.
 
-  Lemma Abs_inv n (a : Tm (S n)) b :
-    rtc RPar'.R (Abs a) b -> exists a', b = Abs a' /\ rtc RPar'.R a a'.
+  Lemma Abs_inv n (a : PTm (S n)) b :
+    rtc RPar'.R (PAbs a) b -> exists a', b = PAbs a' /\ rtc RPar'.R a a'.
   Proof.
-    move E : (Abs a) => b0 h. move : a E.
+    move E : (PAbs a) => b0 h. move : a E.
     elim : b0 b / h.
     - hauto lq:on ctrs:rtc.
     - hauto lq:on ctrs:rtc inv:RPar'.R, rtc.
   Qed.
 
-  Lemma morphing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+  Lemma morphing n m (a b : PTm n) (ρ : fin n -> PTm m) :
     rtc RPar'.R a b ->
-    rtc RPar'.R (subst_Tm ρ a) (subst_Tm ρ b).
+    rtc RPar'.R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof. induction 1; qauto l:on ctrs:rtc use:RPar'.substing. Qed.
 
-  Lemma substing n (a b : Tm (S n)) c :
+  Lemma substing n (a b : PTm (S n)) c :
     rtc RPar'.R a b ->
-    rtc RPar'.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b).
+    rtc RPar'.R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons c VarPTm) b).
   Proof. hauto lq:on use:morphing inv:option. Qed.
 
-  Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
+  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
-    rtc RPar'.R (subst_Tm ρ a) b -> exists b0, rtc RPar'.R a b0 /\ subst_Tm ρ b0 = b.
+    rtc RPar'.R (subst_PTm ρ a) b -> exists b0, rtc RPar'.R a b0 /\ subst_PTm ρ b0 = b.
   Proof.
-    move E  :(subst_Tm ρ a) => u hρ h.
+    move E  :(subst_PTm ρ a) => u hρ h.
     move : a E.
     elim : u b /h.
     - sfirstorder.
@@ -1187,15 +1118,15 @@ Module RPars'.
 
 End RPars'.
 
-Lemma Abs_EPar n a (b : Tm n) :
-  EPar.R (Abs a) b ->
+Lemma Abs_EPar n a (b : PTm n) :
+  EPar.R (PAbs a) b ->
   (exists d, EPar.R a d /\
-     rtc RPar.R (App (ren_Tm shift b) (VarTm var_zero)) d) /\
+     rtc RPar.R (PApp (ren_PTm shift b) (VarPTm var_zero)) d) /\
          (exists d,
          EPar.R a d /\ forall p,
-         rtc RPar.R (Proj p b) (Abs (Proj p d))).
+         rtc RPar.R (PProj p b) (PAbs (PProj p d))).
 Proof.
-  move E : (Abs a) => u h.
+  move E : (PAbs a) => u h.
   move : a E.
   elim : n u b /h => //=.
   - move => n a0 a1 ha iha b ?. subst.
@@ -1216,14 +1147,14 @@ Proof.
   - move => n ? a1 ha iha a0 ?. subst. specialize iha with (1 := eq_refl).
     move : iha => [_ [d [ih0 ih1]]].
     split.
-    + exists (Pair (Proj PL d) (Proj PR d)).
+    + exists (PPair (PProj PL d) (PProj PR d)).
       split; first by apply EPar.PairEta.
       apply : rtc_l.
       apply RPar.AppPair; eauto using RPar.refl.
-      suff h : forall p, rtc RPar.R (App (Proj p (ren_Tm shift a1)) (VarTm var_zero)) (Proj p d) by
+      suff h : forall p, rtc RPar.R (PApp (PProj p (ren_PTm shift a1)) (VarPTm var_zero)) (PProj p d) by
           sfirstorder use:RPars.PairCong.
       move => p. move /(_ p) /RPars.weakening in ih1.
-      apply relations.rtc_transitive with (y := App (ren_Tm shift (Abs (Proj p d))) (VarTm var_zero)).
+      apply relations.rtc_transitive with (y := PApp (ren_PTm shift (PAbs (PProj p d))) (VarPTm var_zero)).
       by eauto using RPars.AppCong, rtc_refl.
       apply relations.rtc_once => /=.
       apply : RPar.AppAbs'; eauto using RPar.refl.
@@ -1239,21 +1170,21 @@ Proof.
       apply rtc_once. apply : RPar.ProjAbs; eauto using RPar.refl.
 Qed.
 
-Lemma Pair_EPar n (a b c : Tm n) :
-  EPar.R (Pair a b) c ->
-  (forall p, exists d, rtc RPar.R (Proj p c) d /\ EPar.R (if p is PL then a else b) d) /\
-    (exists d0 d1, rtc RPar.R (App (ren_Tm shift c) (VarTm var_zero))
-                (Pair (App (ren_Tm shift d0) (VarTm var_zero))(App (ren_Tm shift d1) (VarTm var_zero))) /\
+Lemma Pair_EPar n (a b c : PTm n) :
+  EPar.R (PPair a b) c ->
+  (forall p, exists d, rtc RPar.R (PProj p c) d /\ EPar.R (if p is PL then a else b) d) /\
+    (exists d0 d1, rtc RPar.R (PApp (ren_PTm shift c) (VarPTm var_zero))
+                (PPair (PApp (ren_PTm shift d0) (VarPTm var_zero))(PApp (ren_PTm shift d1) (VarPTm var_zero))) /\
     EPar.R a d0 /\ EPar.R b d1).
 Proof.
-  move E : (Pair a b) => u h. move : a b E.
+  move E : (PPair a b) => u h. move : a b E.
   elim : n u c /h => //=.
   - move => n a0 a1 ha iha a b ?. subst.
     specialize iha with (1 := eq_refl).
     move : iha => [_ [d0 [d1 [ih0 [ih1 ih2]]]]].
     split.
     + move => p.
-      exists (Abs (App (ren_Tm shift (if p is PL then d0 else d1)) (VarTm var_zero))).
+      exists (PAbs (PApp (ren_PTm shift (if p is PL then d0 else d1)) (VarPTm var_zero))).
       split.
       * apply : relations.rtc_transitive.
         ** apply RPars.ProjCong. apply RPars.AbsCong. eassumption.
@@ -1271,7 +1202,7 @@ Proof.
       exists d. split=>//.
       apply : rtc_l. apply RPar.ProjPair; eauto using RPar.refl.
       set q := (X in rtc RPar.R X d).
-      by have -> : q = Proj p a1 by hauto lq:on.
+      by have -> : q = PProj p a1 by hauto lq:on.
     + move  :iha => [iha _].
       move : (iha PL) => [d0 [ih0 ih0']].
       move : (iha PR) => [d1 [ih1 ih1']] {iha}.
@@ -1292,7 +1223,7 @@ Proof.
       split => //.
 Qed.
 
-Lemma commutativity0 n (a b0 b1 : Tm n) :
+Lemma commutativity0 n (a b0 b1 : PTm n) :
   EPar.R a b0 -> RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ EPar.R b1 c.
 Proof.
   move => h. move : b1.
@@ -1300,13 +1231,13 @@ Proof.
   - move => n a b0 ha iha b1 hb.
     move : iha (hb) => /[apply].
     move => [c [ih0 ih1]].
-    exists (Abs (App (ren_Tm shift c) (VarTm var_zero))).
+    exists (PAbs (PApp (ren_PTm shift c) (VarPTm var_zero))).
     split.
     + hauto lq:on ctrs:rtc use:RPars.AbsCong, RPars.AppCong, RPars.renaming.
     + hauto lq:on ctrs:EPar.R use:EPar.refl, EPar.renaming.
   - move => n a b0 hb0 ihb0 b1 /[dup] hb1 {}/ihb0.
     move => [c [ih0 ih1]].
-    exists (Pair (Proj PL c) (Proj PR c)). split.
+    exists (PPair (PProj PL c) (PProj PR c)). split.
     + apply RPars.PairCong;
       by apply RPars.ProjCong.
     + hauto lq:on ctrs:EPar.R use:EPar.refl, EPar.renaming.
@@ -1319,9 +1250,9 @@ Proof.
     elim /RPar.inv => //= _.
     + move =>  a2 a3 b3 b4 h0 h1 [*]. subst.
       move /(_ _ ltac:(by eauto)) : ihb => [b [ihb0 ihb1]].
-      have {}/iha : RPar.R (Abs a2) (Abs a3) by hauto lq:on ctrs:RPar.R.
+      have {}/iha : RPar.R (PAbs a2) (PAbs a3) by hauto lq:on ctrs:RPar.R.
       move => [c [ih0 /Abs_EPar [[d [ih1 ih2]] _]]].
-      exists (subst_Tm (scons b VarTm) d).
+      exists (subst_PTm (scons b VarPTm) d).
       split.
       (* By substitution *)
       * move /RPars.substing  : ih2.
@@ -1337,7 +1268,7 @@ Proof.
       move /Pair_EPar  : ihc1 => [_ [d0 [d1 [ih0 [ih1 ih2]]]]].
       move /RPars.substing : ih0. move /(_ d).
       asimpl => h.
-      exists (Pair (App d0 d) (App d1 d)).
+      exists (PPair (PApp d0 d) (PApp d1 d)).
       split.
       hauto lq:on use:relations.rtc_transitive, RPars.AppCong.
       apply EPar.PairCong; by apply EPar.AppCong.
@@ -1348,7 +1279,7 @@ Proof.
     + move => ? a0 a1 h [*]. subst.
       move /(_ _ ltac:(by eauto using RPar.AbsCong)) : ih0 => [c [ih0 ih1]].
       move /Abs_EPar : ih1 => [_ [d [ih1 ih2]]].
-      exists (Abs (Proj p d)).
+      exists (PAbs (PProj p d)).
       qauto l:on ctrs:EPar.R use:RPars.ProjCong, @relations.rtc_transitive.
     + move => p0 a0 a1 b2 b3 h1 h2 [*]. subst.
       move /(_ _ ltac:(by eauto using RPar.PairCong)) : ih0 => [c [ih0 ih1]].
@@ -1356,13 +1287,12 @@ Proof.
       exists d. split => //.
       hauto lq:on use:RPars.ProjCong, relations.rtc_transitive.
     + hauto lq:on ctrs:EPar.R use:RPars.ProjCong.
-  - hauto lq:on inv:RPar.R ctrs:EPar.R, rtc use:RPars.BindCong.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
   - hauto l:on ctrs:EPar.R inv:RPar.R.
 Qed.
 
-Lemma commutativity1 n (a b0 b1 : Tm n) :
+Lemma commutativity1 n (a b0 b1 : PTm n) :
   EPar.R a b0 -> rtc RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ EPar.R b1 c.
 Proof.
   move => + h. move : b0.
@@ -1371,7 +1301,7 @@ Proof.
   - qauto l:on use:relations.rtc_transitive, commutativity0.
 Qed.
 
-Lemma commutativity n (a b0 b1 : Tm n) :
+Lemma commutativity n (a b0 b1 : PTm n) :
   rtc EPar.R a b0 -> rtc RPar.R a b1 -> exists c, rtc RPar.R b0 c /\ rtc EPar.R b1 c.
   move => h. move : b1. elim : a b0 /h.
   - sfirstorder.
@@ -1380,12 +1310,12 @@ Lemma commutativity n (a b0 b1 : Tm n) :
     hauto q:on ctrs:rtc.
 Qed.
 
-Lemma Abs_EPar' n a (b : Tm n) :
-  EPar.R (Abs a) b ->
+Lemma Abs_EPar' n a (b : PTm n) :
+  EPar.R (PAbs a) b ->
   (exists d, EPar.R a d /\
-          rtc OExp.R (Abs d) b).
+          rtc OExp.R (PAbs d) b).
 Proof.
-  move E : (Abs a) => u h.
+  move E : (PAbs a) => u h.
   move : a E.
   elim : n u b /h => //=.
   - move => n a0 a1 ha iha a ?. subst.
@@ -1397,12 +1327,12 @@ Proof.
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma Proj_EPar' n p a (b : Tm n) :
-  EPar.R (Proj p a) b ->
+Lemma Proj_EPar' n p a (b : PTm n) :
+  EPar.R (PProj p a) b ->
   (exists d, EPar.R a d /\
-          rtc OExp.R (Proj p d) b).
+          rtc OExp.R (PProj p d) b).
 Proof.
-  move E : (Proj p a) => u h.
+  move E : (PProj p a) => u h.
   move : p a E.
   elim : n u b /h => //=.
   - move => n a0 a1 ha iha a p ?. subst.
@@ -1414,11 +1344,11 @@ Proof.
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma App_EPar' n (a b u : Tm n)  :
-  EPar.R (App a b) u ->
-  (exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (App a0 b0) u).
+Lemma App_EPar' n (a b u : PTm n)  :
+  EPar.R (PApp a b) u ->
+  (exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (PApp a0 b0) u).
 Proof.
-  move E : (App a b) => t h.
+  move E : (PApp a b) => t h.
   move : a b E. elim : n t u /h => //=.
   - move => n a0 a1 ha iha a b ?. subst.
     specialize iha with (1 := eq_refl).
@@ -1429,11 +1359,11 @@ Proof.
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma Bind_EPar' n p (a : Tm n) b u  :
-  EPar.R (TBind p a b) u ->
-  (exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (TBind p a0 b0) u).
+Lemma Pair_EPar' n (a b u : PTm n) :
+  EPar.R (PPair a b) u ->
+  exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (PPair a0 b0) u.
 Proof.
-  move E : (TBind p a b) => t h.
+  move E : (PPair a b) => t h.
   move : a b E. elim : n t u /h => //=.
   - move => n a0 a1 ha iha a b ?. subst.
     specialize iha with (1 := eq_refl).
@@ -1444,25 +1374,10 @@ Proof.
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma Pair_EPar' n (a b u : Tm n) :
-  EPar.R (Pair a b) u ->
-  exists a0 b0, EPar.R a a0 /\ EPar.R b b0 /\ rtc OExp.R (Pair a0 b0) u.
-Proof.
-  move E : (Pair a b) => t h.
-  move : a b E. elim : n t u /h => //=.
-  - move => n a0 a1 ha iha a b ?. subst.
-    specialize iha with (1 := eq_refl).
-    hauto lq:on ctrs:OExp.R use:rtc_r.
-  - move => n a0 a1 ha iha a b ?. subst.
-    specialize iha with (1 := eq_refl).
-    hauto lq:on ctrs:OExp.R use:rtc_r.
-  - hauto l:on ctrs:OExp.R.
-Qed.
-
-Lemma Const_EPar' n k (u : Tm n) :
-  EPar.R (Const k) u ->
-  rtc OExp.R (Const k) u.
-  move E : (Const k) => t h.
+Lemma Const_EPar' n k (u : PTm n) :
+  EPar.R (PConst k) u ->
+  rtc OExp.R (PConst k) u.
+  move E : (PConst k) => t h.
   move : k E. elim : n t u /h => //=.
   - move => n a0 a1 h ih k ?. subst.
     specialize ih with (1 := eq_refl).
@@ -1473,10 +1388,10 @@ Lemma Const_EPar' n k (u : Tm n) :
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma Bot_EPar' n (u : Tm n) :
-  EPar.R (Bot) u ->
-  rtc OExp.R (Bot) u.
-  move E : (Bot) => t h.
+Lemma Bot_EPar' n (u : PTm n) :
+  EPar.R (PBot) u ->
+  rtc OExp.R (PBot) u.
+  move E : (PBot) => t h.
   move : E. elim : n t u /h => //=.
   - move => n a0 a1 h ih ?. subst.
     specialize ih with (1 := eq_refl).
@@ -1487,10 +1402,10 @@ Lemma Bot_EPar' n (u : Tm n) :
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma Univ_EPar' n i (u : Tm n) :
-  EPar.R (Univ i) u ->
-  rtc OExp.R (Univ i) u.
-  move E : (Univ i) => t h.
+Lemma Univ_EPar' n i (u : PTm n) :
+  EPar.R (PUniv i) u ->
+  rtc OExp.R (PUniv i) u.
+  move E : (PUniv i) => t h.
   move : E. elim : n t u /h => //=.
   - move => n a0 a1 h ih ?. subst.
     specialize ih with (1 := eq_refl).
@@ -1501,14 +1416,14 @@ Lemma Univ_EPar' n i (u : Tm n) :
   - hauto l:on ctrs:OExp.R.
 Qed.
 
-Lemma EPar_diamond n (c a1 b1 : Tm n) :
+Lemma EPar_diamond n (c a1 b1 : PTm n) :
   EPar.R c a1 ->
   EPar.R c b1 ->
   exists d2, EPar.R a1 d2 /\ EPar.R b1 d2.
 Proof.
   move => h. move : b1. elim : n c a1 / h.
   - move => n c a1 ha iha b1 /iha [d2 [hd0 hd1]].
-    exists(Abs (App (ren_Tm shift d2) (VarTm var_zero))).
+    exists(PAbs (PApp (ren_PTm shift d2) (VarPTm var_zero))).
     hauto lq:on ctrs:EPar.R use:EPar.renaming.
   - hauto lq:on rew:off ctrs:EPar.R.
   - hauto lq:on use:EPar.refl.
@@ -1516,62 +1431,54 @@ Proof.
     move /Abs_EPar' => [d [hd0 hd1]].
     move : iha hd0; repeat move/[apply].
     move => [d2 [h0 h1]].
-    have : EPar.R (Abs d) (Abs d2) by eauto using EPar.AbsCong.
+    have : EPar.R (PAbs d) (PAbs d2) by eauto using EPar.AbsCong.
     move : OExp.commutativity0 hd1; repeat move/[apply].
     move => [d1 [hd1 hd2]].
     exists d1. hauto lq:on ctrs:EPar.R use:OExp.merge.
   - move => n a0 a1 b0 b1 ha iha hb ihb c.
     move /App_EPar' => [a2][b2][/iha [a3 h0]][/ihb [b3 h1]]h2 {iha ihb}.
-    have : EPar.R (App a2 b2)(App a3 b3)
+    have : EPar.R (PApp a2 b2)(PApp a3 b3)
       by hauto l:on use:EPar.AppCong.
     move : OExp.commutativity0 h2; repeat move/[apply].
     move => [d h].
     exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
   - move => n a0 a1 b0 b1 ha iha hb ihb c.
     move /Pair_EPar' => [a2][b2][/iha [a3 h0]][/ihb [b3 h1]]h2 {iha ihb}.
-    have : EPar.R (Pair a2 b2)(Pair a3 b3)
+    have : EPar.R (PPair a2 b2)(PPair a3 b3)
       by hauto l:on use:EPar.PairCong.
     move : OExp.commutativity0 h2; repeat move/[apply].
     move => [d h].
     exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
   - move => n p a0 a1 ha iha b.
     move /Proj_EPar' => [d [/iha [d2 h] h1]] {iha}.
-    have : EPar.R (Proj p d) (Proj p d2)
+    have : EPar.R (PProj p d) (PProj p d2)
       by hauto l:on use:EPar.ProjCong.
     move : OExp.commutativity0 h1; repeat move/[apply].
     move => [d1 h1].
     exists d1. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
-  - move => n p a0 a1 b0 b1 ha iha hb ihb c.
-    move /Bind_EPar' => [a2][b2][/iha [a3 h0]][/ihb [b3 h1]]h2 {iha ihb}.
-    have : EPar.R (TBind p a2 b2)(TBind p a3 b3)
-      by hauto l:on use:EPar.BindCong.
-    move : OExp.commutativity0 h2; repeat move/[apply].
-    move => [d h].
-    exists d. hauto lq:on rew:off ctrs:EPar.R use:OExp.merge.
   - qauto use:Const_EPar', EPar.refl.
   - qauto use:Univ_EPar', EPar.refl.
   - qauto use:Bot_EPar', EPar.refl.
 Qed.
 
-Function tstar {n} (a : Tm n) :=
+Function tstar {n} (a : PTm n) :=
   match a with
-  | VarTm i => a
-  | Abs a => Abs (tstar a)
-  | App (Abs a) b => subst_Tm (scons (tstar b) VarTm) (tstar a)
-  | App (Pair a b) c =>
-      Pair (App (tstar a) (tstar c)) (App (tstar b) (tstar c))
-  | App a b => App (tstar a) (tstar b)
-  | Pair a b => Pair (tstar a) (tstar b)
-  | Proj p (Pair a b) => if p is PL then (tstar a) else (tstar b)
-  | Proj p (Abs a) => (Abs (Proj p (tstar a)))
-  | Proj p a => Proj p (tstar a)
-  | TBind p a b => TBind p (tstar a) (tstar b)
-  | Const k => Const k
-  | Univ i => Univ i
-  | Bot => Bot
+  | VarPTm i => a
+  | PAbs a => PAbs (tstar a)
+  | PApp (PAbs a) b => subst_PTm (scons (tstar b) VarPTm) (tstar a)
+  | PApp (PPair a b) c =>
+      PPair (PApp (tstar a) (tstar c)) (PApp (tstar b) (tstar c))
+  | PApp a b => PApp (tstar a) (tstar b)
+  | PPair a b => PPair (tstar a) (tstar b)
+  | PProj p (PPair a b) => if p is PL then (tstar a) else (tstar b)
+  | PProj p (PAbs a) => (PAbs (PProj p (tstar a)))
+  | PProj p a => PProj p (tstar a)
+  | PConst k => PConst k
+  | PUniv i => PUniv i
+  | PBot => PBot
   end.
 
-Lemma RPar_triangle n (a : Tm n) : forall b, RPar.R a b -> RPar.R b (tstar a).
+Lemma RPar_triangle n (a : PTm n) : forall b, RPar.R a b -> RPar.R b (tstar a).
 Proof.
   apply tstar_ind => {n a}.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
@@ -1587,25 +1494,23 @@ Proof.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
   - hauto lq:on inv:RPar.R ctrs:RPar.R.
-  - hauto lq:on inv:RPar.R ctrs:RPar.R.
 Qed.
 
-Function tstar' {n} (a : Tm n) :=
+Function tstar' {n} (a : PTm n) :=
   match a with
-  | VarTm i => a
-  | Abs a => Abs (tstar' a)
-  | App (Abs a) b => subst_Tm (scons (tstar' b) VarTm) (tstar' a)
-  | App a b => App (tstar' a) (tstar' b)
-  | Pair a b => Pair (tstar' a) (tstar' b)
-  | Proj p (Pair a b) => if p is PL then (tstar' a) else (tstar' b)
-  | Proj p a => Proj p (tstar' a)
-  | TBind p a b => TBind p (tstar' a) (tstar' b)
-  | Const k => Const k
-  | Univ i => Univ i
-  | Bot => Bot
+  | VarPTm i => a
+  | PAbs a => PAbs (tstar' a)
+  | PApp (PAbs a) b => subst_PTm (scons (tstar' b) VarPTm) (tstar' a)
+  | PApp a b => PApp (tstar' a) (tstar' b)
+  | PPair a b => PPair (tstar' a) (tstar' b)
+  | PProj p (PPair a b) => if p is PL then (tstar' a) else (tstar' b)
+  | PProj p a => PProj p (tstar' a)
+  | PConst k => PConst k
+  | PUniv i => PUniv i
+  | PBot => PBot
   end.
 
-Lemma RPar'_triangle n (a : Tm n) : forall b, RPar'.R a b -> RPar'.R b (tstar' a).
+Lemma RPar'_triangle n (a : PTm n) : forall b, RPar'.R a b -> RPar'.R b (tstar' a).
 Proof.
   apply tstar'_ind => {n a}.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
@@ -1619,22 +1524,21 @@ Proof.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
   - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
-  - hauto lq:on inv:RPar'.R ctrs:RPar'.R.
 Qed.
 
-Lemma RPar_diamond n (c a1 b1 : Tm n) :
+Lemma RPar_diamond n (c a1 b1 : PTm n) :
   RPar.R c a1 ->
   RPar.R c b1 ->
   exists d2, RPar.R a1 d2 /\ RPar.R b1 d2.
 Proof. hauto l:on use:RPar_triangle. Qed.
 
-Lemma RPar'_diamond n (c a1 b1 : Tm n) :
+Lemma RPar'_diamond n (c a1 b1 : PTm n) :
   RPar'.R c a1 ->
   RPar'.R c b1 ->
   exists d2, RPar'.R a1 d2 /\ RPar'.R b1 d2.
 Proof. hauto l:on use:RPar'_triangle. Qed.
 
-Lemma RPar_confluent n (c a1 b1 : Tm n) :
+Lemma RPar_confluent n (c a1 b1 : PTm n) :
   rtc RPar.R c a1 ->
   rtc RPar.R c b1 ->
   exists d2, rtc RPar.R a1 d2 /\ rtc RPar.R b1 d2.
@@ -1642,7 +1546,7 @@ Proof.
   sfirstorder use:relations.diamond_confluent, RPar_diamond.
 Qed.
 
-Lemma EPar_confluent n (c a1 b1 : Tm n) :
+Lemma EPar_confluent n (c a1 b1 : PTm n) :
   rtc EPar.R c a1 ->
   rtc EPar.R c b1 ->
   exists d2, rtc EPar.R a1 d2 /\ rtc EPar.R b1 d2.
@@ -1650,52 +1554,35 @@ Proof.
   sfirstorder use:relations.diamond_confluent, EPar_diamond.
 Qed.
 
-Definition prov_bind {n} p0 A0 B0 (a : Tm n)  :=
-  match a with
-  | TBind p A B => p = p0 /\ rtc Par.R A A0 /\ rtc Par.R B B0
-  | _ => False
-  end.
-
-Definition prov_univ {n} i0 (a : Tm n) :=
-  match a with
-  | Univ i => i = i0
-  | _ => False
-  end.
-
-
-Inductive prov {n} : Tm n -> Tm n -> Prop :=
-| P_Bind p A A0 B B0  :
-  rtc Par.R A A0 ->
-  rtc Par.R B B0 ->
-  prov (TBind p A B) (TBind p A0 B0)
+Inductive prov {n} : PTm n -> PTm n -> Prop :=
 | P_Abs h a :
-  (forall b, prov h (subst_Tm (scons b VarTm) a)) ->
-  prov h (Abs a)
+  (forall b, prov h (subst_PTm (scons b VarPTm) a)) ->
+  prov h (PAbs a)
 | P_App h a b  :
   prov h a ->
-  prov h (App a b)
+  prov h (PApp a b)
 | P_Pair h a b :
   prov h a ->
   prov h b ->
-  prov h (Pair a b)
+  prov h (PPair a b)
 | P_Proj h p a :
   prov h a ->
-  prov h (Proj p a)
+  prov h (PProj p a)
 | P_Const k  :
-  prov (Const k) (Const k)
+  prov (PConst k) (PConst k)
 | P_Var i :
-  prov (VarTm i) (VarTm i)
+  prov (VarPTm i) (VarPTm i)
 | P_Univ i :
-  prov (Univ i) (Univ i)
+  prov (PUniv i) (PUniv i)
 | P_Bot :
-  prov Bot Bot.
+  prov PBot PBot.
 
-Lemma ERed_EPar n (a b : Tm n) : ERed.R a b -> EPar.R a b.
+Lemma ERed_EPar n (a b : PTm n) : ERed.R a b -> EPar.R a b.
 Proof.
   induction 1; hauto lq:on ctrs:EPar.R use:EPar.refl.
 Qed.
 
-Lemma EPar_ERed n (a b : Tm n) : EPar.R a b -> rtc ERed.R a b.
+Lemma EPar_ERed n (a b : PTm n) : EPar.R a b -> rtc ERed.R a b.
 Proof.
   move => h. elim : n a b /h.
   - eauto using rtc_r, ERed.AppEta.
@@ -1705,57 +1592,56 @@ Proof.
   - eauto using EReds.AppCong.
   - eauto using EReds.PairCong.
   - eauto using EReds.ProjCong.
-  - eauto using EReds.BindCong.
   - auto using rtc_refl.
   - auto using rtc_refl.
   - auto using rtc_refl.
 Qed.
 
-Lemma EPar_Par n (a b : Tm n) : EPar.R a b -> Par.R a b.
+Lemma EPar_Par n (a b : PTm n) : EPar.R a b -> Par.R a b.
 Proof.
   move => h. elim : n a b /h; qauto ctrs:Par.R.
 Qed.
 
-Lemma RPar_Par n (a b : Tm n) : RPar.R a b -> Par.R a b.
+Lemma RPar_Par n (a b : PTm n) : RPar.R a b -> Par.R a b.
 Proof.
   move => h. elim : n a b /h; hauto lq:on ctrs:Par.R.
 Qed.
 
-Lemma rtc_idem n (R : Tm n -> Tm n -> Prop) (a b : Tm n) : rtc (rtc R) a b -> rtc R a b.
+Lemma rtc_idem n (R : PTm n -> PTm n -> Prop) (a b : PTm n) : rtc (rtc R) a b -> rtc R a b.
 Proof.
   induction 1; hauto l:on use:@relations.rtc_transitive, @rtc_r.
 Qed.
 
-Lemma EPars_EReds {n} (a b : Tm n) : rtc EPar.R a b <-> rtc ERed.R a b.
+Lemma EPars_EReds {n} (a b : PTm n) : rtc EPar.R a b <-> rtc ERed.R a b.
 Proof.
   sfirstorder use:@relations.rtc_subrel, EPar_ERed, rtc_idem, ERed_EPar.
 Qed.
 
-Lemma prov_rpar n (u : Tm n) a b : prov u a -> RPar.R a b -> prov u b.
+Lemma prov_rpar n (u : PTm n) a b : prov u a -> RPar.R a b -> prov u b.
 Proof.
   move => h.
   move : b.
   elim : u a / h.
-  - qauto l:on ctrs:prov inv:RPar.R use:@rtc_r, RPar_Par.
+  (* - qauto l:on ctrs:prov inv:RPar.R use:@rtc_r, RPar_Par. *)
   - hauto lq:on ctrs:prov inv:RPar.R use:RPar.substing.
   - move => h a b ha iha b0.
     elim /RPar.inv => //= _.
     + move => a0 a1 b1 b2 h0 h1 [*]. subst.
-      have {}iha :  prov h (Abs a1) by hauto lq:on ctrs:RPar.R.
+      have {}iha :  prov h (PAbs a1) by hauto lq:on ctrs:RPar.R.
       hauto lq:on inv:prov use:RPar.substing.
     + move => a0 a1 b1 b2 c0 c1.
       move => h0 h1 h2 [*]. subst.
-      have {}iha : prov h (Pair a1 b2) by hauto lq:on ctrs:RPar.R.
+      have {}iha : prov h (PPair a1 b2) by hauto lq:on ctrs:RPar.R.
       hauto lq:on inv:prov ctrs:prov.
     + hauto lq:on ctrs:prov.
   - hauto lq:on ctrs:prov inv:RPar.R.
   - move => h p a ha iha b.
     elim /RPar.inv => //= _.
     + move => p0 a0 a1 h0 [*]. subst.
-      have {iha} :  prov h (Abs a1) by hauto lq:on ctrs:RPar.R.
+      have {iha} :  prov h (PAbs a1) by hauto lq:on ctrs:RPar.R.
       hauto lq:on ctrs:prov inv:prov use:RPar.substing.
     + move => p0 a0 a1 b0 b1 h0 h1 [*]. subst.
-      have {iha} : prov h (Pair a1 b1) by hauto lq:on ctrs:RPar.R.
+      have {iha} : prov h (PPair a1 b1) by hauto lq:on ctrs:RPar.R.
       qauto l:on inv:prov.
     + hauto lq:on ctrs:prov.
   - hauto lq:on ctrs:prov inv:RPar.R.
@@ -1765,33 +1651,23 @@ Proof.
 Qed.
 
 
-Lemma prov_lam n (u : Tm n) a : prov u a <-> prov u (Abs (App (ren_Tm shift a) (VarTm var_zero))).
+Lemma prov_lam n (u : PTm n) a : prov u a <-> prov u (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))).
 Proof.
   split.
   move => h. constructor. move => b. asimpl. by constructor.
   inversion 1; subst.
-  specialize H2 with (b := Const TPi).
+  specialize H2 with (b := PBot).
   move : H2. asimpl. inversion 1; subst. done.
 Qed.
 
-Lemma prov_pair n (u : Tm n) a : prov u a <-> prov u (Pair (Proj PL a) (Proj PR a)).
+Lemma prov_pair n (u : PTm n) a : prov u a <-> prov u (PPair (PProj PL a) (PProj PR a)).
 Proof. hauto lq:on inv:prov ctrs:prov. Qed.
 
-Lemma prov_ered n (u : Tm n) a b : prov u a -> ERed.R a b -> prov u b.
+Lemma prov_ered n (u : PTm n) a b : prov u a -> ERed.R a b -> prov u b.
 Proof.
   move => h.
   move : b.
   elim : u a / h.
-  - move => p A A0 B B0 hA hB b.
-    elim /ERed.inv => // _.
-    + move => a0 *. subst.
-      rewrite -prov_lam.
-      by constructor.
-    + move => a0 *. subst.
-      rewrite -prov_pair.
-      by constructor.
-    + qauto l:on ctrs:prov use:@rtc_r, ERed_EPar, EPar_Par.
-    + qauto l:on ctrs:prov use:@rtc_r, ERed_EPar, EPar_Par.
   - move => h a ha iha b.
     elim /ERed.inv => // _.
     + move => a0 *. subst.
@@ -1819,26 +1695,25 @@ Proof.
   - hauto lq:on inv:ERed.R, prov ctrs:prov.
 Qed.
 
-Lemma prov_ereds n (u : Tm n) a b : prov u a -> rtc ERed.R a b -> prov u b.
+Lemma prov_ereds n (u : PTm n) a b : prov u a -> rtc ERed.R a b -> prov u b.
 Proof.
   induction 2; sfirstorder use:prov_ered.
 Qed.
 
-Fixpoint extract {n} (a : Tm n) : Tm n :=
+Fixpoint extract {n} (a : PTm n) : PTm n :=
   match a with
-  | TBind p A B => TBind p A B
-  | Abs a => subst_Tm (scons Bot VarTm) (extract a)
-  | App a b => extract a
-  | Pair a b => extract a
-  | Proj p a => extract a
-  | Const k => Const k
-  | VarTm i => VarTm i
-  | Univ i => Univ i
-  | Bot => Bot
+  | PAbs a => subst_PTm (scons PBot VarPTm) (extract a)
+  | PApp a b => extract a
+  | PPair a b => extract a
+  | PProj p a => extract a
+  | PConst k => PConst k
+  | VarPTm i => VarPTm i
+  | PUniv i => PUniv i
+  | PBot => PBot
   end.
 
-Lemma ren_extract n m (a : Tm n) (ξ : fin n -> fin m) :
-  extract (ren_Tm ξ a) = ren_Tm ξ (extract a).
+Lemma ren_extract n m (a : PTm n) (ξ : fin n -> fin m) :
+  extract (ren_PTm ξ a) = ren_PTm ξ (extract a).
 Proof.
   move : m ξ. elim : n/a.
   - sfirstorder.
@@ -1851,12 +1726,11 @@ Proof.
   - hauto q:on.
   - sfirstorder.
   - sfirstorder.
-  - sfirstorder.
 Qed.
 
-Lemma ren_morphing n m (a : Tm n) (ρ : fin n -> Tm m) :
+Lemma ren_morphing n m (a : PTm n) (ρ : fin n -> PTm m) :
   (forall i, ρ i = extract (ρ i)) ->
-  extract (subst_Tm ρ a) = subst_Tm ρ (extract a).
+  extract (subst_PTm ρ a) = subst_PTm ρ (extract a).
 Proof.
   move : m ρ.
   elim : n /a => n //=.
@@ -1869,45 +1743,38 @@ Proof.
   - by asimpl.
 Qed.
 
-Lemma ren_subst_bot n (a : Tm (S n)) :
-  extract (subst_Tm (scons Bot VarTm) a) = subst_Tm (scons Bot VarTm) (extract a).
+Lemma ren_subst_bot n (a : PTm (S n)) :
+  extract (subst_PTm (scons PBot VarPTm) a) = subst_PTm (scons PBot VarPTm) (extract a).
 Proof.
   apply ren_morphing. destruct i as [i|] => //=.
 Qed.
 
-Definition prov_extract_spec {n} u (a : Tm n) :=
+Definition prov_extract_spec {n} u (a : PTm n) :=
   match u with
-  | TBind p A B => exists A0 B0, extract a = TBind p A0 B0 /\ rtc Par.R A A0 /\ rtc Par.R B B0
-  | Univ i => extract a = Univ i
-  | VarTm i => extract a = VarTm i
-  | (Const i) => extract a = (Const i)
-  | Bot => extract a = Bot
+  | PUniv i => extract a = PUniv i
+  | VarPTm i => extract a = VarPTm i
+  | (PConst i) => extract a = (PConst i)
+  | PBot => extract a = PBot
   | _ => True
   end.
 
-Lemma prov_extract n u (a : Tm n) :
+Lemma prov_extract n u (a : PTm n) :
   prov u a -> prov_extract_spec u a.
 Proof.
   move => h.
   elim : u a /h.
-  - sfirstorder.
   - move => h a ha ih.
     case : h ha ih => //=.
     + move => i ha ih.
-      move /(_ Bot) in ih.
+      move /(_ PBot) in ih.
       rewrite -ih.
       by rewrite ren_subst_bot.
-    + move => p A B h ih.
-      move /(_ Bot) : ih => [A0][B0][h0][h1]h2.
-      rewrite ren_subst_bot in h0.
-      rewrite h0.
-      eauto.
-    + move => p _ /(_ Bot).
+    + move => p _ /(_ PBot).
       by rewrite ren_subst_bot.
-    + move => i h /(_ Bot).
+    + move => i h /(_ PBot).
       by rewrite ren_subst_bot => ->.
-    + move /(_ Bot).
-      move => h /(_ Bot).
+    + move /(_ PBot).
+      move => h /(_ PBot).
       by rewrite ren_subst_bot.
   - hauto lq:on.
   - hauto lq:on.
@@ -1922,21 +1789,21 @@ Definition union {A : Type} (R0 R1 : A -> A -> Prop) a b :=
   R0 a b \/ R1 a b.
 
 Module ERPar.
-  Definition R {n} (a b : Tm n) := union RPar.R EPar.R a b.
-  Lemma RPar {n} (a b : Tm n) : RPar.R a b -> R a b.
+  Definition R {n} (a b : PTm n) := union RPar.R EPar.R a b.
+  Lemma RPar {n} (a b : PTm n) : RPar.R a b -> R a b.
   Proof. sfirstorder. Qed.
 
-  Lemma EPar {n} (a b : Tm n) : EPar.R a b -> R a b.
+  Lemma EPar {n} (a b : PTm n) : EPar.R a b -> R a b.
   Proof. sfirstorder. Qed.
 
-  Lemma refl {n} ( a : Tm n) : ERPar.R a a.
+  Lemma refl {n} ( a : PTm n) : ERPar.R a a.
   Proof.
     sfirstorder use:RPar.refl, EPar.refl.
   Qed.
 
-  Lemma ProjCong n p (a0 a1 : Tm n) :
+  Lemma ProjCong n p (a0 a1 : PTm n) :
     R a0 a1 ->
-    rtc R (Proj p a0) (Proj p a1).
+    rtc R (PProj p a0) (PProj p a1).
   Proof.
     move => [].
     - move => h.
@@ -1949,9 +1816,9 @@ Module ERPar.
       by apply EPar.ProjCong.
   Qed.
 
-  Lemma AbsCong n (a0 a1 : Tm (S n)) :
+  Lemma AbsCong n (a0 a1 : PTm (S n)) :
     R a0 a1 ->
-    rtc R (Abs a0) (Abs a1).
+    rtc R (PAbs a0) (PAbs a1).
   Proof.
     move => [].
     - move => h.
@@ -1964,10 +1831,10 @@ Module ERPar.
       by apply EPar.AbsCong.
   Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     R a0 a1 ->
     R b0 b1 ->
-    rtc R (App a0 b0) (App a1 b1).
+    rtc R (PApp a0 b0) (PApp a1 b1).
   Proof.
     move => [] + [].
     - sfirstorder use:RPar.AppCong, @rtc_once.
@@ -1984,30 +1851,10 @@ Module ERPar.
     - sfirstorder use:EPar.AppCong, @rtc_once.
   Qed.
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1:
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     R a0 a1 ->
     R b0 b1 ->
-    rtc R (TBind p a0 b0) (TBind p a1 b1).
-  Proof.
-    move => [] + [].
-    - sfirstorder use:RPar.BindCong, @rtc_once.
-    - move => h0 h1.
-      apply : rtc_l.
-      left. apply RPar.BindCong; eauto; apply RPar.refl.
-      apply rtc_once.
-      hauto l:on use:EPar.BindCong, EPar.refl.
-    - move => h0 h1.
-      apply : rtc_l.
-      left. apply RPar.BindCong; eauto; apply RPar.refl.
-      apply rtc_once.
-      hauto l:on use:EPar.BindCong, EPar.refl.
-    - sfirstorder use:EPar.BindCong, @rtc_once.
-  Qed.
-
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
-    R a0 a1 ->
-    R b0 b1 ->
-    rtc R (Pair a0 b0) (Pair a1 b1).
+    rtc R (PPair a0 b0) (PPair a1 b1).
   Proof.
     move => [] + [].
     - sfirstorder use:RPar.PairCong, @rtc_once.
@@ -2024,15 +1871,15 @@ Module ERPar.
     - sfirstorder use:EPar.PairCong, @rtc_once.
   Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     sfirstorder use:EPar.renaming, RPar.renaming.
   Qed.
 
 End ERPar.
 
-Hint Resolve ERPar.AppCong ERPar.refl ERPar.AbsCong ERPar.PairCong ERPar.ProjCong ERPar.BindCong : erpar.
+Hint Resolve ERPar.AppCong ERPar.refl ERPar.AbsCong ERPar.PairCong ERPar.ProjCong : erpar.
 
 Module ERPars.
   #[local]Ltac solve_s_rec :=
@@ -2041,37 +1888,31 @@ Module ERPars.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     rtc ERPar.R a0 a1 ->
     rtc ERPar.R b0 b1 ->
-    rtc ERPar.R (App a0 b0) (App a1 b1).
+    rtc ERPar.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma AbsCong n (a0 a1 : Tm (S n)) :
+  Lemma AbsCong n (a0 a1 : PTm (S n)) :
     rtc ERPar.R a0 a1 ->
-    rtc ERPar.R (Abs a0) (Abs a1).
+    rtc ERPar.R (PAbs a0) (PAbs a1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     rtc ERPar.R a0 a1 ->
     rtc ERPar.R b0 b1 ->
-    rtc ERPar.R (Pair a0 b0) (Pair a1 b1).
+    rtc ERPar.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1 : Tm n) :
+  Lemma ProjCong n p (a0 a1 : PTm n) :
     rtc ERPar.R a0 a1 ->
-    rtc ERPar.R (Proj p a0) (Proj p a1).
+    rtc ERPar.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (a0 a1 : Tm n) b0 b1:
+  Lemma renaming n (a0 a1 : PTm n) m (ξ : fin n -> fin m) :
     rtc ERPar.R a0 a1 ->
-    rtc ERPar.R b0 b1 ->
-    rtc ERPar.R (TBind p a0 b0) (TBind p a1 b1).
-  Proof. solve_s. Qed.
-
-  Lemma renaming n (a0 a1 : Tm n) m (ξ : fin n -> fin m) :
-    rtc ERPar.R a0 a1 ->
-    rtc ERPar.R (ren_Tm ξ a0) (ren_Tm ξ a1).
+    rtc ERPar.R (ren_PTm ξ a0) (ren_PTm ξ a1).
   Proof.
     induction 1.
     - apply rtc_refl.
@@ -2080,16 +1921,16 @@ Module ERPars.
 
 End ERPars.
 
-Lemma ERPar_Par n (a b : Tm n) : ERPar.R a b -> Par.R a b.
+Lemma ERPar_Par n (a b : PTm n) : ERPar.R a b -> Par.R a b.
 Proof.
   sfirstorder use:EPar_Par, RPar_Par.
 Qed.
 
-Lemma Par_ERPar n (a b : Tm n) : Par.R a b -> rtc ERPar.R a b.
+Lemma Par_ERPar n (a b : PTm n) : Par.R a b -> rtc ERPar.R a b.
 Proof.
   move => h. elim : n a b /h.
   - move => n a0 a1 b0 b1 ha iha hb ihb.
-    suff ? : rtc ERPar.R (App (Abs a0) b0) (App (Abs a1) b1).
+    suff ? : rtc ERPar.R (PApp (PAbs a0) b0) (PApp (PAbs a1) b1).
     apply : relations.rtc_transitive; eauto.
     apply rtc_once. apply ERPar.RPar.
     by apply RPar.AppAbs; eauto using RPar.refl.
@@ -2116,30 +1957,29 @@ Proof.
   - sfirstorder use:ERPars.AppCong.
   - sfirstorder use:ERPars.PairCong.
   - sfirstorder use:ERPars.ProjCong.
-  - sfirstorder use:ERPars.BindCong.
   - sfirstorder.
   - sfirstorder.
   - sfirstorder.
 Qed.
 
-Lemma Pars_ERPar n (a b : Tm n) : rtc Par.R a b -> rtc ERPar.R a b.
+Lemma Pars_ERPar n (a b : PTm n) : rtc Par.R a b -> rtc ERPar.R a b.
 Proof.
   induction 1; hauto l:on use:Par_ERPar, @relations.rtc_transitive.
 Qed.
 
-Lemma Par_ERPar_iff n (a b : Tm n) : rtc Par.R a b <-> rtc ERPar.R a b.
+Lemma Par_ERPar_iff n (a b : PTm n) : rtc Par.R a b <-> rtc ERPar.R a b.
 Proof.
   split.
   sfirstorder use:Pars_ERPar, @relations.rtc_subrel.
   sfirstorder use:ERPar_Par, @relations.rtc_subrel.
 Qed.
 
-Lemma RPar_ERPar n (a b : Tm n) : rtc RPar.R a b -> rtc ERPar.R a b.
+Lemma RPar_ERPar n (a b : PTm n) : rtc RPar.R a b -> rtc ERPar.R a b.
 Proof.
   sfirstorder use:@relations.rtc_subrel.
 Qed.
 
-Lemma EPar_ERPar n (a b : Tm n) : rtc EPar.R a b -> rtc ERPar.R a b.
+Lemma EPar_ERPar n (a b : PTm n) : rtc EPar.R a b -> rtc ERPar.R a b.
 Proof.
   sfirstorder use:@relations.rtc_subrel.
 Qed.
@@ -2205,7 +2045,7 @@ Module HindleyRosenFacts (M : HindleyRosen).
 End HindleyRosenFacts.
 
 Module HindleyRosenER <: HindleyRosen.
-  Definition A := Tm.
+  Definition A := PTm.
   Definition R0 n := rtc (@RPar.R n).
   Definition R1 n := rtc (@EPar.R n).
   Lemma diamond_R0 : forall n, relations.diamond (R0 n).
@@ -2223,7 +2063,7 @@ End HindleyRosenER.
 
 Module ERFacts := HindleyRosenFacts HindleyRosenER.
 
-Lemma rtc_union n (a b : Tm n) :
+Lemma rtc_union n (a b : PTm n) :
   rtc (union RPar.R EPar.R) a b <->
     rtc (union (rtc RPar.R) (rtc EPar.R)) a b.
 Proof.
@@ -2245,7 +2085,7 @@ Proof.
       sfirstorder.
 Qed.
 
-Lemma prov_erpar n (u : Tm n) a b : prov u a -> ERPar.R a b -> prov u b.
+Lemma prov_erpar n (u : PTm n) a b : prov u a -> ERPar.R a b -> prov u b.
 Proof.
   move => h [].
   - sfirstorder use:prov_rpar.
@@ -2253,7 +2093,7 @@ Proof.
     sfirstorder use:prov_ereds.
 Qed.
 
-Lemma prov_pars n (u : Tm n) a b : prov u a -> rtc Par.R a b -> prov u b.
+Lemma prov_pars n (u : PTm n) a b : prov u a -> rtc Par.R a b -> prov u b.
 Proof.
   move => h /Pars_ERPar.
   move => h0.
@@ -2263,15 +2103,15 @@ Proof.
   - hauto lq:on use:prov_erpar.
 Qed.
 
-Lemma Par_confluent n (a b c : Tm n) :
+Lemma Par_confluent n (a b c : PTm n) :
   rtc Par.R a b ->
   rtc Par.R a c ->
   exists d, rtc Par.R b d /\ rtc Par.R c d.
 Proof.
   move : n a b c.
-  suff : forall (n : nat) (a b c : Tm n),
+  suff : forall (n : nat) (a b c : PTm n),
       rtc ERPar.R a b ->
-      rtc ERPar.R a c -> exists d : Tm n, rtc ERPar.R b d /\ rtc ERPar.R c d.
+      rtc ERPar.R a c -> exists d : PTm n, rtc ERPar.R b d /\ rtc ERPar.R c d.
   move => h n a b c h0 h1.
   apply Par_ERPar_iff in h0, h1.
   move : h h0 h1; repeat move/[apply].
@@ -2282,32 +2122,32 @@ Proof.
   specialize h with (n := n).
   rewrite /HindleyRosenER.A in h.
   rewrite /ERPar.R.
-  have eq : (fun a0 b0 : Tm n => union RPar.R EPar.R a0 b0) = union RPar.R EPar.R by reflexivity.
+  have eq : (fun a0 b0 : PTm n => union RPar.R EPar.R a0 b0) = union RPar.R EPar.R by reflexivity.
   rewrite !{}eq.
   move /rtc_union => + /rtc_union.
   move : h; repeat move/[apply].
   hauto lq:on use:rtc_union.
 Qed.
 
-Lemma pars_univ_inv n i (c : Tm n) :
-  rtc Par.R (Univ i) c ->
-  extract c = Univ i.
+Lemma pars_univ_inv n i (c : PTm n) :
+  rtc Par.R (PUniv i) c ->
+  extract c = PUniv i.
 Proof.
-  have : prov (Univ i) (Univ i : Tm n) by sfirstorder.
+  have : prov (PUniv i) (Univ i : PTm n) by sfirstorder.
   move : prov_pars. repeat move/[apply].
   apply prov_extract.
 Qed.
 
-Lemma pars_const_inv n i (c : Tm n) :
+Lemma pars_const_inv n i (c : PTm n) :
   rtc Par.R (Const i) c ->
   extract c = Const i.
 Proof.
-  have : prov (Const i) (Const i : Tm n) by sfirstorder.
+  have : prov (Const i) (Const i : PTm n) by sfirstorder.
   move : prov_pars. repeat move/[apply].
   apply prov_extract.
 Qed.
 
-Lemma pars_pi_inv n p (A : Tm n) B C :
+Lemma pars_pi_inv n p (A : PTm n) B C :
   rtc Par.R (TBind p A B) C ->
   exists A0 B0, extract C = TBind p A0 B0 /\
              rtc Par.R A A0 /\ rtc Par.R B B0.
@@ -2318,15 +2158,15 @@ Proof.
 Qed.
 
 Lemma pars_var_inv n (i : fin n) C :
-  rtc Par.R (VarTm i) C ->
-  extract C = VarTm i.
+  rtc Par.R (VarPTm i) C ->
+  extract C = VarPTm i.
 Proof.
-  have : prov (VarTm i) (VarTm i) by hauto lq:on ctrs:prov, rtc.
+  have : prov (VarPTm i) (VarPTm i) by hauto lq:on ctrs:prov, rtc.
   move : prov_pars. repeat move/[apply].
   apply prov_extract.
 Qed.
 
-Lemma pars_univ_inj n i j (C : Tm n) :
+Lemma pars_univ_inj n i j (C : PTm n) :
   rtc Par.R (Univ i) C ->
   rtc Par.R (Univ j) C ->
   i = j.
@@ -2334,7 +2174,7 @@ Proof.
   sauto l:on use:pars_univ_inv.
 Qed.
 
-Lemma pars_const_inj n i j (C : Tm n) :
+Lemma pars_const_inj n i j (C : PTm n) :
   rtc Par.R (Const i) C ->
   rtc Par.R (Const j) C ->
   i = j.
@@ -2342,7 +2182,7 @@ Proof.
   sauto l:on use:pars_const_inv.
 Qed.
 
-Lemma pars_pi_inj n p0 p1 (A0 A1 : Tm n) B0 B1 C :
+Lemma pars_pi_inj n p0 p1 (A0 A1 : PTm n) B0 B1 C :
   rtc Par.R (TBind p0 A0 B0) C ->
   rtc Par.R (TBind p1 A1 B1) C ->
   exists A2 B2, p1 = p0 /\ rtc Par.R A0 A2 /\ rtc Par.R A1 A2 /\
@@ -2353,10 +2193,10 @@ Proof.
   exists A2, B2. hauto l:on.
 Qed.
 
-Definition join {n} (a b : Tm n) :=
+Definition join {n} (a b : PTm n) :=
   exists c, rtc Par.R a c /\ rtc Par.R b c.
 
-Lemma join_transitive n (a b c : Tm n) :
+Lemma join_transitive n (a b c : PTm n) :
   join a b -> join b c -> join a c.
 Proof.
   rewrite /join.
@@ -2366,26 +2206,26 @@ Proof.
   eauto using relations.rtc_transitive.
 Qed.
 
-Lemma join_symmetric n (a b : Tm n) :
+Lemma join_symmetric n (a b : PTm n) :
   join a b -> join b a.
 Proof. sfirstorder unfold:join. Qed.
 
-Lemma join_refl n (a : Tm n) : join a a.
+Lemma join_refl n (a : PTm n) : join a a.
 Proof. hauto lq:on ctrs:rtc unfold:join. Qed.
 
 Lemma join_univ_inj n i j :
-  join (Univ i : Tm n) (Univ j) -> i = j.
+  join (Univ i : PTm n) (Univ j) -> i = j.
 Proof.
   sfirstorder use:pars_univ_inj.
 Qed.
 
 Lemma join_const_inj n i j :
-  join (Const i : Tm n) (Const j) -> i = j.
+  join (Const i : PTm n) (Const j) -> i = j.
 Proof.
   sfirstorder use:pars_const_inj.
 Qed.
 
-Lemma join_pi_inj n p0 p1 (A0 A1 : Tm n) B0 B1 :
+Lemma join_pi_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
   join (TBind p0 A0 B0) (TBind p1 A1 B1) ->
   p0 = p1 /\ join A0 A1 /\ join B0 B1.
 Proof.
@@ -2394,14 +2234,14 @@ Proof.
   sfirstorder unfold:join.
 Qed.
 
-Lemma join_substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
+Lemma join_substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
   join a b ->
-  join (subst_Tm ρ a) (subst_Tm ρ b).
+  join (subst_PTm ρ a) (subst_PTm ρ b).
 Proof. hauto lq:on unfold:join use:Pars.substing. Qed.
 
-Fixpoint ne {n} (a : Tm n) :=
+Fixpoint ne {n} (a : PTm n) :=
   match a with
-  | VarTm i => true
+  | VarPTm i => true
   | TBind _ A B => false
   | App a b => ne a && nf b
   | Abs a => false
@@ -2411,9 +2251,9 @@ Fixpoint ne {n} (a : Tm n) :=
   | Const _ => false
   | Bot => true
   end
-with nf {n} (a : Tm n) :=
+with nf {n} (a : PTm n) :=
   match a with
-  | VarTm i => true
+  | VarPTm i => true
   | TBind _ A B => nf A && nf B
   | App a b => ne a && nf b
   | Abs a => nf a
@@ -2427,8 +2267,8 @@ end.
 Lemma ne_nf n a : @ne n a -> nf a.
 Proof. elim : a => //=. Qed.
 
-Definition wn {n} (a : Tm n) := exists b, rtc RPar'.R a b /\ nf b.
-Definition wne {n} (a : Tm n) := exists b, rtc RPar'.R a b /\ ne b.
+Definition wn {n} (a : PTm n) := exists b, rtc RPar'.R a b /\ nf b.
+Definition wne {n} (a : PTm n) := exists b, rtc RPar'.R a b /\ ne b.
 
 (* Weakly neutral implies weakly normal *)
 Lemma wne_wn n a : @wne n a -> wn a.
@@ -2438,18 +2278,18 @@ Proof. sfirstorder use:ne_nf. Qed.
 Lemma nf_wn n v : @nf n v -> wn v.
 Proof. sfirstorder ctrs:rtc. Qed.
 
-Lemma nf_refl n (a b : Tm n) (h : RPar'.R a b) : (nf a -> b = a) /\ (ne a -> b = a).
+Lemma nf_refl n (a b : PTm n) (h : RPar'.R a b) : (nf a -> b = a) /\ (ne a -> b = a).
 Proof.
   elim : a b /h => //=; solve [hauto b:on].
 Qed.
 
-Lemma ne_nf_ren n m (a : Tm n) (ξ : fin n -> fin m) :
-  (ne a <-> ne (ren_Tm ξ a)) /\ (nf a <-> nf (ren_Tm ξ a)).
+Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
+  (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)).
 Proof.
   move : m ξ. elim : n / a => //=; solve [hauto b:on].
 Qed.
 
-Lemma wne_app n (a b : Tm n) :
+Lemma wne_app n (a b : PTm n) :
   wne a -> wn b -> wne (App a b).
 Proof.
   move => [a0 [? ?]] [b0 [? ?]].
@@ -2470,14 +2310,14 @@ Proof.
   hauto lqb:on use:RPars'.BindCong.
 Qed.
 
-Lemma wn_pair n (a b : Tm n) : wn a -> wn b -> wn (Pair a b).
+Lemma wn_pair n (a b : PTm n) : wn a -> wn b -> wn (Pair a b).
 Proof.
   move => [a0 [? ?]] [b0 [? ?]].
   exists (Pair a0 b0).
   hauto lqb:on use:RPars'.PairCong.
 Qed.
 
-Lemma wne_proj n p (a : Tm n) : wne a -> wne (Proj p a).
+Lemma wne_proj n p (a : PTm n) : wne a -> wne (Proj p a).
 Proof.
   move => [a0 [? ?]].
   exists (Proj p a0). hauto lqb:on use:RPars'.ProjCong.
@@ -2486,17 +2326,17 @@ Qed.
 Create HintDb nfne.
 #[export]Hint Resolve nf_wn ne_nf wne_wn nf_refl : nfne.
 
-Lemma ne_nf_antiren n m (a : Tm n) (ρ : fin n -> Tm m) :
+Lemma ne_nf_antiren n m (a : PTm n) (ρ : fin n -> PTm m) :
   (forall i, var_or_const (ρ i)) ->
-  (ne (subst_Tm ρ a) -> ne a) /\ (nf (subst_Tm ρ a) -> nf a).
+  (ne (subst_PTm ρ a) -> ne a) /\ (nf (subst_PTm ρ a) -> nf a).
 Proof.
   move : m ρ. elim : n / a => //;
    hauto b:on drew:off use:RPar.var_or_const_up.
 Qed.
 
-Lemma wn_antirenaming n m a (ρ : fin n -> Tm m) :
+Lemma wn_antirenaming n m a (ρ : fin n -> PTm m) :
   (forall i, var_or_const (ρ i)) ->
-  wn (subst_Tm ρ a) -> wn a.
+  wn (subst_PTm ρ a) -> wn a.
 Proof.
   rewrite /wn => hρ.
   move => [v [rv nfv]].
@@ -2507,7 +2347,7 @@ Proof.
   by eapply ne_nf_antiren.
 Qed.
 
-Lemma ext_wn n (a : Tm n) :
+Lemma ext_wn n (a : PTm n) :
     wn (App a Bot) ->
     wn a.
 Proof.
@@ -2515,7 +2355,7 @@ Proof.
   move : a E.
   move : hv.
   elim : a0 v / hr.
-  - hauto q:on inv:Tm ctrs:rtc b:on db: nfne.
+  - hauto q:on inv:PTm ctrs:rtc b:on db: nfne.
   - move => a0 a1 a2 hr0 hr1 ih hnfa2.
     move /(_ hnfa2) in ih.
     move => a.
@@ -2524,7 +2364,7 @@ Proof.
     + move => a0 a3 b2 b3 ? ? [? ?] ? [? ?]. subst.
       have ? : b3 = (Bot) by hauto lq:on inv:RPar'.R. subst.
       suff : wn (Abs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
-      have : wn (subst_Tm (scons (Bot) VarTm) a3) by sfirstorder.
+      have : wn (subst_PTm (scons (Bot) VarPTm) a3) by sfirstorder.
       move => h. apply wn_abs.
       move : h. apply wn_antirenaming.
       hauto lq:on rew:off inv:option.
@@ -2532,41 +2372,41 @@ Proof.
 Qed.
 
 Module Join.
-  Lemma ProjCong p n (a0 a1 : Tm n) :
+  Lemma ProjCong p n (a0 a1 : PTm n) :
     join a0 a1 ->
     join (Proj p a0) (Proj p a1).
   Proof. hauto lq:on use:Pars.ProjCong unfold:join. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     join a0 a1 ->
     join b0 b1 ->
     join (Pair a0 b0) (Pair a1 b1).
   Proof. hauto lq:on use:Pars.PairCong unfold:join. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : Tm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     join a0 a1 ->
     join b0 b1 ->
     join (App a0 b0) (App a1 b1).
   Proof. hauto lq:on use:Pars.AppCong. Qed.
 
-  Lemma AbsCong n (a b : Tm (S n)) :
+  Lemma AbsCong n (a b : PTm (S n)) :
     join a b ->
     join (Abs a) (Abs b).
   Proof. hauto lq:on use:Pars.AbsCong. Qed.
 
-  Lemma renaming n m (a b : Tm n) (ξ : fin n -> fin m) :
-    join a b -> join (ren_Tm ξ a) (ren_Tm ξ b).
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    join a b -> join (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     induction 1; hauto lq:on use:Pars.renaming.
   Qed.
 
-  Lemma weakening n (a b : Tm n)  :
-    join a b -> join (ren_Tm shift a) (ren_Tm shift b).
+  Lemma weakening n (a b : PTm n)  :
+    join a b -> join (ren_PTm shift a) (ren_PTm shift b).
   Proof.
     apply renaming.
   Qed.
 
-  Lemma FromPar n (a b : Tm n) :
+  Lemma FromPar n (a b : PTm n) :
     Par.R a b ->
     join a b.
   Proof.
@@ -2574,12 +2414,12 @@ Module Join.
   Qed.
 End Join.
 
-Lemma abs_eq n a (b : Tm n) :
-  join (Abs a) b <-> join a (App (ren_Tm shift b) (VarTm var_zero)).
+Lemma abs_eq n a (b : PTm n) :
+  join (Abs a) b <-> join a (App (ren_PTm shift b) (VarPTm var_zero)).
 Proof.
   split.
   - move => /Join.weakening h.
-    have {h} : join (App (ren_Tm shift (Abs a)) (VarTm var_zero)) (App (ren_Tm shift b) (VarTm var_zero))
+    have {h} : join (App (ren_PTm shift (Abs a)) (VarPTm var_zero)) (App (ren_PTm shift b) (VarPTm var_zero))
       by hauto l:on use:Join.AppCong, join_refl.
     simpl.
     move => ?. apply : join_transitive; eauto.
@@ -2590,7 +2430,7 @@ Proof.
     apply join_symmetric. apply Join.FromPar. apply Par.AppEta. apply Par.refl.
 Qed.
 
-Lemma pair_eq n (a0 a1 b : Tm n) :
+Lemma pair_eq n (a0 a1 b : PTm n) :
   join (Pair a0 a1) b <-> join a0 (Proj PL b) /\ join a1 (Proj PR b).
 Proof.
   split.
@@ -2606,7 +2446,7 @@ Proof.
     apply Join.FromPar. hauto lq:on ctrs:Par.R use:Par.refl.
 Qed.
 
-Lemma join_pair_inj n (a0 a1 b0 b1 : Tm n) :
+Lemma join_pair_inj n (a0 a1 b0 b1 : PTm n) :
   join (Pair a0 a1) (Pair b0 b1) <-> join a0 b0 /\ join a1 b1.
 Proof.
   split; last by hauto lq:on use:Join.PairCong.

From 398a18d7709e8dc9ff0933c9ceb8dc5a0ed1bd3f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 24 Jan 2025 14:58:35 -0700
Subject: [PATCH 12/13] Finish renaming

---
 theories/fp_red.v | 135 +++++++++++++++++-----------------------------
 1 file changed, 48 insertions(+), 87 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3482f45..bbe58d9 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2133,26 +2133,16 @@ Lemma pars_univ_inv n i (c : PTm n) :
   rtc Par.R (PUniv i) c ->
   extract c = PUniv i.
 Proof.
-  have : prov (PUniv i) (Univ i : PTm n) by sfirstorder.
+  have : prov (PUniv i) (PUniv i : PTm n) by sfirstorder.
   move : prov_pars. repeat move/[apply].
   apply prov_extract.
 Qed.
 
 Lemma pars_const_inv n i (c : PTm n) :
-  rtc Par.R (Const i) c ->
-  extract c = Const i.
+  rtc Par.R (PConst i) c ->
+  extract c = PConst i.
 Proof.
-  have : prov (Const i) (Const i : PTm n) by sfirstorder.
-  move : prov_pars. repeat move/[apply].
-  apply prov_extract.
-Qed.
-
-Lemma pars_pi_inv n p (A : PTm n) B C :
-  rtc Par.R (TBind p A B) C ->
-  exists A0 B0, extract C = TBind p A0 B0 /\
-             rtc Par.R A A0 /\ rtc Par.R B B0.
-Proof.
-  have : prov (TBind p A B) (TBind p A B) by hauto lq:on ctrs:prov, rtc.
+  have : prov (PConst i) (PConst i : PTm n) by sfirstorder.
   move : prov_pars. repeat move/[apply].
   apply prov_extract.
 Qed.
@@ -2167,32 +2157,21 @@ Proof.
 Qed.
 
 Lemma pars_univ_inj n i j (C : PTm n) :
-  rtc Par.R (Univ i) C ->
-  rtc Par.R (Univ j) C ->
+  rtc Par.R (PUniv i) C ->
+  rtc Par.R (PUniv j) C ->
   i = j.
 Proof.
   sauto l:on use:pars_univ_inv.
 Qed.
 
 Lemma pars_const_inj n i j (C : PTm n) :
-  rtc Par.R (Const i) C ->
-  rtc Par.R (Const j) C ->
+  rtc Par.R (PConst i) C ->
+  rtc Par.R (PConst j) C ->
   i = j.
 Proof.
   sauto l:on use:pars_const_inv.
 Qed.
 
-Lemma pars_pi_inj n p0 p1 (A0 A1 : PTm n) B0 B1 C :
-  rtc Par.R (TBind p0 A0 B0) C ->
-  rtc Par.R (TBind p1 A1 B1) C ->
-  exists A2 B2, p1 = p0 /\ rtc Par.R A0 A2 /\ rtc Par.R A1 A2 /\
-             rtc Par.R B0 B2 /\ rtc Par.R B1 B2.
-Proof.
-  move /pars_pi_inv => [A2 [B2 [? [h0 h1]]]].
-  move /pars_pi_inv => [A3 [B3 [? [h2 h3]]]].
-  exists A2, B2. hauto l:on.
-Qed.
-
 Definition join {n} (a b : PTm n) :=
   exists c, rtc Par.R a c /\ rtc Par.R b c.
 
@@ -2214,26 +2193,17 @@ Lemma join_refl n (a : PTm n) : join a a.
 Proof. hauto lq:on ctrs:rtc unfold:join. Qed.
 
 Lemma join_univ_inj n i j :
-  join (Univ i : PTm n) (Univ j) -> i = j.
+  join (PUniv i : PTm n) (PUniv j) -> i = j.
 Proof.
   sfirstorder use:pars_univ_inj.
 Qed.
 
 Lemma join_const_inj n i j :
-  join (Const i : PTm n) (Const j) -> i = j.
+  join (PConst i : PTm n) (PConst j) -> i = j.
 Proof.
   sfirstorder use:pars_const_inj.
 Qed.
 
-Lemma join_pi_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
-  join (TBind p0 A0 B0) (TBind p1 A1 B1) ->
-  p0 = p1 /\ join A0 A1 /\ join B0 B1.
-Proof.
-  move => [c []].
-  move : pars_pi_inj; repeat move/[apply].
-  sfirstorder unfold:join.
-Qed.
-
 Lemma join_substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
   join a b ->
   join (subst_PTm ρ a) (subst_PTm ρ b).
@@ -2242,26 +2212,24 @@ Proof. hauto lq:on unfold:join use:Pars.substing. Qed.
 Fixpoint ne {n} (a : PTm n) :=
   match a with
   | VarPTm i => true
-  | TBind _ A B => false
-  | App a b => ne a && nf b
-  | Abs a => false
-  | Univ _ => false
-  | Proj _ a => ne a
-  | Pair _ _ => false
-  | Const _ => false
-  | Bot => true
+  | PApp a b => ne a && nf b
+  | PAbs a => false
+  | PUniv _ => false
+  | PProj _ a => ne a
+  | PPair _ _ => false
+  | PConst _ => false
+  | PBot => true
   end
 with nf {n} (a : PTm n) :=
   match a with
   | VarPTm i => true
-  | TBind _ A B => nf A && nf B
-  | App a b => ne a && nf b
-  | Abs a => nf a
-  | Univ _ => true
-  | Proj _ a => ne a
-  | Pair a b => nf a && nf b
-  | Const _ => true
-  | Bot => true
+  | PApp a b => ne a && nf b
+  | PAbs a => nf a
+  | PUniv _ => true
+  | PProj _ a => ne a
+  | PPair a b => nf a && nf b
+  | PConst _ => true
+  | PBot => true
 end.
 
 Lemma ne_nf n a : @ne n a -> nf a.
@@ -2290,37 +2258,30 @@ Proof.
 Qed.
 
 Lemma wne_app n (a b : PTm n) :
-  wne a -> wn b -> wne (App a b).
+  wne a -> wn b -> wne (PApp a b).
 Proof.
   move => [a0 [? ?]] [b0 [? ?]].
-  exists (App a0 b0). hauto b:on drew:off use:RPars'.AppCong.
+  exists (PApp a0 b0). hauto b:on drew:off use:RPars'.AppCong.
 Qed.
 
-Lemma wn_abs n a (h : wn a) : @wn n (Abs a).
+Lemma wn_abs n a (h : wn a) : @wn n (PAbs a).
 Proof.
   move : h => [v [? ?]].
-  exists (Abs v).
+  exists (PAbs v).
   eauto using RPars'.AbsCong.
 Qed.
 
-Lemma wn_bind n p A B : wn A -> wn B -> wn (@TBind n p A B).
-Proof.
-  move => [A0 [? ?]] [B0 [? ?]].
-  exists (TBind p A0 B0).
-  hauto lqb:on use:RPars'.BindCong.
-Qed.
-
-Lemma wn_pair n (a b : PTm n) : wn a -> wn b -> wn (Pair a b).
+Lemma wn_pair n (a b : PTm n) : wn a -> wn b -> wn (PPair a b).
 Proof.
   move => [a0 [? ?]] [b0 [? ?]].
-  exists (Pair a0 b0).
+  exists (PPair a0 b0).
   hauto lqb:on use:RPars'.PairCong.
 Qed.
 
-Lemma wne_proj n p (a : PTm n) : wne a -> wne (Proj p a).
+Lemma wne_proj n p (a : PTm n) : wne a -> wne (PProj p a).
 Proof.
   move => [a0 [? ?]].
-  exists (Proj p a0). hauto lqb:on use:RPars'.ProjCong.
+  exists (PProj p a0). hauto lqb:on use:RPars'.ProjCong.
 Qed.
 
 Create HintDb nfne.
@@ -2348,10 +2309,10 @@ Proof.
 Qed.
 
 Lemma ext_wn n (a : PTm n) :
-    wn (App a Bot) ->
+    wn (PApp a PBot) ->
     wn a.
 Proof.
-  move E : (App a (Bot)) => a0 [v [hr hv]].
+  move E : (PApp a (PBot)) => a0 [v [hr hv]].
   move : a E.
   move : hv.
   elim : a0 v / hr.
@@ -2362,9 +2323,9 @@ Proof.
     case : a0 hr0=>// => b0 b1.
     elim /RPar'.inv=>// _.
     + move => a0 a3 b2 b3 ? ? [? ?] ? [? ?]. subst.
-      have ? : b3 = (Bot) by hauto lq:on inv:RPar'.R. subst.
-      suff : wn (Abs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
-      have : wn (subst_PTm (scons (Bot) VarPTm) a3) by sfirstorder.
+      have ? : b3 = (PBot) by hauto lq:on inv:RPar'.R. subst.
+      suff : wn (PAbs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
+      have : wn (subst_PTm (scons (PBot) VarPTm) a3) by sfirstorder.
       move => h. apply wn_abs.
       move : h. apply wn_antirenaming.
       hauto lq:on rew:off inv:option.
@@ -2374,24 +2335,24 @@ Qed.
 Module Join.
   Lemma ProjCong p n (a0 a1 : PTm n) :
     join a0 a1 ->
-    join (Proj p a0) (Proj p a1).
+    join (PProj p a0) (PProj p a1).
   Proof. hauto lq:on use:Pars.ProjCong unfold:join. Qed.
 
   Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
     join a0 a1 ->
     join b0 b1 ->
-    join (Pair a0 b0) (Pair a1 b1).
+    join (PPair a0 b0) (PPair a1 b1).
   Proof. hauto lq:on use:Pars.PairCong unfold:join. Qed.
 
   Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
     join a0 a1 ->
     join b0 b1 ->
-    join (App a0 b0) (App a1 b1).
+    join (PApp a0 b0) (PApp a1 b1).
   Proof. hauto lq:on use:Pars.AppCong. Qed.
 
   Lemma AbsCong n (a b : PTm (S n)) :
     join a b ->
-    join (Abs a) (Abs b).
+    join (PAbs a) (PAbs b).
   Proof. hauto lq:on use:Pars.AbsCong. Qed.
 
   Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
@@ -2415,11 +2376,11 @@ Module Join.
 End Join.
 
 Lemma abs_eq n a (b : PTm n) :
-  join (Abs a) b <-> join a (App (ren_PTm shift b) (VarPTm var_zero)).
+  join (PAbs a) b <-> join a (PApp (ren_PTm shift b) (VarPTm var_zero)).
 Proof.
   split.
   - move => /Join.weakening h.
-    have {h} : join (App (ren_PTm shift (Abs a)) (VarPTm var_zero)) (App (ren_PTm shift b) (VarPTm var_zero))
+    have {h} : join (PApp (ren_PTm shift (PAbs a)) (VarPTm var_zero)) (PApp (ren_PTm shift b) (VarPTm var_zero))
       by hauto l:on use:Join.AppCong, join_refl.
     simpl.
     move => ?. apply : join_transitive; eauto.
@@ -2431,12 +2392,12 @@ Proof.
 Qed.
 
 Lemma pair_eq n (a0 a1 b : PTm n) :
-  join (Pair a0 a1) b <-> join a0 (Proj PL b) /\ join a1 (Proj PR b).
+  join (PPair a0 a1) b <-> join a0 (PProj PL b) /\ join a1 (PProj PR b).
 Proof.
   split.
   - move => h.
     have /Join.ProjCong {}h := h.
-    have h0 : forall p, join (if p is PL then a0 else a1) (Proj p (Pair a0 a1))
+    have h0 : forall p, join (if p is PL then a0 else a1) (PProj p (PPair a0 a1))
                      by hauto lq:on use:join_symmetric, Join.FromPar, Par.ProjPair', Par.refl.
     hauto lq:on rew:off use:join_transitive, join_symmetric.
   - move => [h0 h1].
@@ -2447,11 +2408,11 @@ Proof.
 Qed.
 
 Lemma join_pair_inj n (a0 a1 b0 b1 : PTm n) :
-  join (Pair a0 a1) (Pair b0 b1) <-> join a0 b0 /\ join a1 b1.
+  join (PPair a0 a1) (PPair b0 b1) <-> join a0 b0 /\ join a1 b1.
 Proof.
   split; last by hauto lq:on use:Join.PairCong.
   move /pair_eq => [h0 h1].
-  have : join (Proj PL (Pair b0 b1)) b0 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
-  have : join (Proj PR (Pair b0 b1)) b1 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
+  have : join (PProj PL (PPair b0 b1)) b0 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
+  have : join (PProj PR (PPair b0 b1)) b1 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
   eauto using join_transitive.
 Qed.

From 9c17ec5cac4f702a69146c86cc244669a82b739e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 1 Apr 2025 23:21:39 -0400
Subject: [PATCH 13/13] Start refactoring to unscoped syntax

---
 Makefile                       |    8 +-
 theories/Autosubst2/fintype.v  |  419 ----------
 theories/Autosubst2/syntax.v   | 1362 ++++++++++++--------------------
 theories/Autosubst2/unscoped.v |  213 +++++
 theories/fp_red.v              |   84 +-
 theories/typing.v              |  251 ++++++
 6 files changed, 994 insertions(+), 1343 deletions(-)
 delete mode 100644 theories/Autosubst2/fintype.v
 create mode 100644 theories/Autosubst2/unscoped.v
 create mode 100644 theories/typing.v

diff --git a/Makefile b/Makefile
index ef5b76f..c56e771 100644
--- a/Makefile
+++ b/Makefile
@@ -6,7 +6,7 @@ theories: $(COQMAKEFILE) FORCE
 validate: $(COQMAKEFILE) FORCE
 	$(MAKE) -f $(COQMAKEFILE) validate
 
-$(COQMAKEFILE) : theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v
+$(COQMAKEFILE) : theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/unscoped.v
 	$(COQBIN)coq_makefile -f _CoqProject -o $(COQMAKEFILE)
 
 install: $(COQMAKEFILE)
@@ -15,13 +15,13 @@ install: $(COQMAKEFILE)
 uninstall: $(COQMAKEFILE)
 	$(MAKE) -f $(COQMAKEFILE) uninstall
 
-theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v : syntax.sig
-	autosubst -f -v ge813 -s coq -o theories/Autosubst2/syntax.v syntax.sig
+theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/unscoped.v : syntax.sig
+	autosubst -f -v ge813 -s ucoq -o theories/Autosubst2/syntax.v syntax.sig
 
 .PHONY: clean FORCE
 
 clean:
 	test ! -f $(COQMAKEFILE) || ( $(MAKE) -f $(COQMAKEFILE) clean && rm $(COQMAKEFILE) )
-	rm -f theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v
+	rm -f theories/Autosubst2/syntax.v theories/Autosubst2/core.v theories/Autosubst2/fintype.v theories/Autosubst2/unscoped.v
 
 FORCE:
diff --git a/theories/Autosubst2/fintype.v b/theories/Autosubst2/fintype.v
deleted file mode 100644
index 99508b6..0000000
--- a/theories/Autosubst2/fintype.v
+++ /dev/null
@@ -1,419 +0,0 @@
-(** * Autosubst Header for Scoped Syntax
-    Our development utilises well-scoped de Bruijn syntax. This means that the de Bruijn indices are taken from finite types. As a consequence, any kind of substitution or environment used in conjunction with well-scoped syntax takes the form of a mapping from some finite type _I^n_. In particular, _renamings_ are mappings _I^n -> I^m_. Here we develop the theory of how these parts interact.
-
-Version: December 11, 2019.
-*)
-Require Import core.
-Require Import Setoid Morphisms Relation_Definitions.
-
-Set Implicit Arguments.
-Unset Strict Implicit.
-
-Definition ap {X Y} (f : X -> Y) {x y : X} (p : x = y) : f x = f y :=
-  match p with eq_refl => eq_refl end.
-
-Definition apc {X Y} {f g : X -> Y} {x y : X} (p : f = g) (q : x = y) : f x = g y :=
-  match q with eq_refl => match p with eq_refl => eq_refl end end.
-
-(** ** Primitives of the Sigma Calculus
-    We implement the finite type with _n_ elements, _I^n_, as the _n_-fold iteration of the Option Type. _I^0_ is implemented as the empty type.
-*)
-
-Fixpoint fin (n : nat) : Type :=
-  match n with
-  | 0 => False
-  | S m => option (fin m)
-  end.
-
-(** Renamings and Injective Renamings
-     _Renamings_ are mappings between finite types.
-*)
-Definition ren (m n : nat) : Type := fin m -> fin n.
-
-Definition id {X} := @Datatypes.id X.
-
-Definition idren {k: nat} : ren k k := @Datatypes.id (fin k).
-
-(** We give a special name, to the newest element in a non-empty finite type, as it usually corresponds to a freshly bound variable. *)
-Definition var_zero {n : nat} : fin (S n) := None.
-
-Definition null {T} (i : fin 0) : T := match i with end.
-
-Definition shift {n : nat} : ren n (S n) :=
-  Some.
-
-(** Extension of Finite Mappings
-    Assume we are given a mapping _f_ from _I^n_ to some type _X_, then we can _extend_ this mapping with a new value from _x : X_ to a mapping from _I^n+1_ to _X_. We denote this operation by _x . f_ and define it as follows:
-*)
-Definition scons {X : Type} {n : nat} (x : X) (f : fin n -> X) (m : fin (S n)) : X :=
-  match m with
-  | None => x
-  | Some i => f i
-  end.
-
-#[ export ]
-Hint Opaque scons : rewrite.
-
-(** ** Type Class Instances for Notation *)
-
-(** *** Type classes for renamings. *)
-
-Class Ren1 (X1  : Type) (Y Z : Type) :=
-  ren1 : X1 -> Y -> Z.
-
-Class Ren2 (X1 X2 : Type) (Y Z : Type) :=
-  ren2 : X1 -> X2 -> Y -> Z.
-
-Class Ren3 (X1 X2 X3 : Type) (Y Z : Type) :=
-  ren3 : X1 -> X2 -> X3 -> Y -> Z.
-
-Class Ren4 (X1 X2 X3 X4 : Type) (Y Z : Type) :=
-  ren4 : X1 -> X2 -> X3 -> X4 -> Y -> Z.
-
-Class Ren5 (X1 X2 X3 X4 X5 : Type) (Y Z : Type) :=
-  ren5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z.
-
-Module RenNotations.
-  Notation "s ⟨ xi1 ⟩" := (ren1  xi1 s) (at level 7, left associativity, format "s  ⟨ xi1 ⟩") : subst_scope.
-
-  Notation "s ⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ⟩") : subst_scope.
-
-  Notation "s ⟨ xi1 ; xi2 ; xi3 ⟩" := (ren3 xi1 xi2 xi3 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ; xi3 ⟩") : subst_scope.
-
-  Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩" := (ren4  xi1 xi2 xi3 xi4 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩") : subst_scope.
-
-  Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩" := (ren5  xi1 xi2 xi3 xi4 xi5 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩") : subst_scope.
-
-  Notation "⟨ xi ⟩" := (ren1 xi) (at level 1, left associativity, format "⟨ xi ⟩") : fscope.
-
-  Notation "⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2) (at level 1, left associativity, format "⟨ xi1 ; xi2 ⟩") : fscope.
-End RenNotations.
-
-(** *** Type Classes for Substiution *)
-
-Class Subst1 (X1 : Type) (Y Z: Type) :=
-  subst1 : X1 -> Y -> Z.
-
-Class Subst2 (X1 X2 : Type) (Y Z: Type) :=
-  subst2 : X1 -> X2 -> Y  -> Z.
-
-Class Subst3 (X1 X2 X3 : Type) (Y Z: Type) :=
-  subst3 : X1 -> X2 -> X3 ->  Y  -> Z.
-
-Class Subst4 (X1 X2 X3 X4: Type) (Y Z: Type) :=
-  subst4 : X1 -> X2 -> X3 -> X4 -> Y  -> Z.
-
-Class Subst5 (X1 X2 X3 X4 X5 : Type) (Y Z: Type) :=
-  subst5 : X1 -> X2 -> X3 -> X4 -> X5  -> Y  -> Z.
-
-Module SubstNotations.
-  Notation "s [ sigma ]" := (subst1 sigma s) (at level 7, left associativity, format "s '/' [ sigma ]") : subst_scope.
-
-  Notation "s [ sigma ; tau ]" := (subst2 sigma tau s) (at level 7, left associativity, format "s '/' [ sigma ; '/'  tau ]") : subst_scope.
-End SubstNotations.
-
-(** ** Type Class for Variables *)
-Class Var X Y :=
-  ids : X -> Y.
-
-
-(** ** Proofs for substitution primitives *)
-
-(** Forward Function Composition
-    Substitutions represented as functions are ubiquitious in this development and we often have to compose them, without talking about their pointwise behaviour.
-    That is, we are interested in the forward compostion of functions, _f o g_, for which we introduce a convenient notation, "f >> g". The direction of the arrow serves as a reminder of the _forward_ nature of this composition, that is first apply _f_, then _g_. *)
-
-Arguments funcomp {X Y Z} (g)%fscope (f)%fscope.
-
-Module CombineNotations.
-  Notation "f >> g" := (funcomp g f) (at level 50) : fscope.
-
-  Notation "s .: sigma" := (scons s sigma) (at level 55, sigma at next level, right associativity) : subst_scope.
-
-  #[ global ]
-  Open Scope fscope.
-  #[ global ]
-  Open Scope subst_scope.
-End CombineNotations.
-
-Import CombineNotations.
-
-
-(** Generic lifting operation for renamings *)
-Definition up_ren {m n} (xi : ren m n) : ren (S m) (S n) :=
-  var_zero .: xi >> shift.
-
-(** Generic proof that lifting of renamings composes. *)
-Lemma up_ren_ren k l m (xi: ren k l) (zeta : ren l m) (rho: ren k m) (E: forall x, (xi >> zeta) x = rho x) :
-  forall x, (up_ren xi >> up_ren zeta) x = up_ren rho x.
-Proof.
-  intros [x|].
-  - cbn. unfold funcomp. now rewrite <- E.
-  - reflexivity.
-Qed.
-
-Arguments up_ren_ren {k l m} xi zeta rho E.
-
-Lemma fin_eta {X} (f g : fin 0 -> X) :
-  pointwise_relation _ eq f g.
-Proof. intros []. Qed.
-
-(** Eta laws *)
-Lemma scons_eta' {T} {n : nat} (f : fin (S n) -> T) :
-  pointwise_relation _ eq (f var_zero .: (funcomp f shift)) f.
-Proof. intros x. destruct x; reflexivity. Qed.
-
-Lemma scons_eta_id' {n : nat} :
-  pointwise_relation (fin (S n)) eq (var_zero .: shift) id.
-Proof. intros x. destruct x; reflexivity. Qed.
-
-Lemma scons_comp' {T:Type} {U} {m} (s: T) (sigma: fin m -> T) (tau: T -> U) :
-  pointwise_relation _ eq (funcomp tau (s .: sigma)) ((tau s) .: (funcomp tau sigma)).
-Proof. intros x. destruct x; reflexivity. Qed.
-
-(* Lemma scons_tail'_pointwise {X} {n} (s : X) (f : fin n -> X) : *)
-(*   pointwise_relation _ eq (funcomp (scons s f) shift) f. *)
-(* Proof. *)
-(*   reflexivity. *)
-(* Qed. *)
-
-(* Lemma scons_tail' {X} {n} (s : X) (f : fin n -> X) x : *)
-(*   (scons s f (shift x)) = f x. *)
-(* Proof. *)
-(*   reflexivity. *)
-(* Qed. *)
-
-(* Morphism for Setoid Rewriting. The only morphism that can be defined statically. *)
-#[export] Instance scons_morphism {X: Type} {n:nat} :
-  Proper (eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons X n).
-Proof.
-  intros t t' -> sigma tau H.
-  intros [x|].
-  cbn. apply H.
-  reflexivity.
-Qed.
-
-#[export] Instance scons_morphism2 {X: Type} {n: nat} :
-  Proper (eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons X n).
-Proof.
-  intros ? t -> sigma tau H ? x ->.
-  destruct x as [x|].
-  cbn. apply H.
-  reflexivity.
-Qed.
-
-(** ** Variadic Substitution Primitives *)
-
-Fixpoint shift_p (p : nat) {n} : ren n (p + n) :=
-  fun n => match p with
-        | 0 => n
-        | S p => Some (shift_p p n)
-        end.
-
-Fixpoint scons_p {X: Type} {m : nat} : forall {n} (f : fin m -> X) (g : fin n -> X), fin (m + n)  -> X.
-Proof.
-  destruct m.
-  - intros n f g. exact g.
-  - intros n f g. cbn. apply scons.
-    + exact (f var_zero).
-    + apply scons_p.
-      * intros z. exact (f (Some z)).
-      * exact g.
-Defined.
-
-#[export] Hint Opaque scons_p : rewrite.
-
-#[export] Instance scons_p_morphism {X: Type} {m n:nat} :
-  Proper (pointwise_relation _ eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons_p X m n).
-Proof.
-  intros sigma sigma' Hsigma tau tau' Htau.
-  intros x.
-  induction m.
-  - cbn. apply Htau.
-  - cbn. change (fin (S m + n)) with (fin (S (m + n))) in x.
-    destruct x as [x|].
-    + cbn. apply IHm.
-      intros ?. apply Hsigma.
-    + cbn. apply Hsigma.
-Qed.
-
-#[export] Instance scons_p_morphism2 {X: Type} {m n:nat} :
-  Proper (pointwise_relation _ eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons_p X m n).
-Proof.
-  intros sigma sigma' Hsigma tau tau' Htau ? x ->.
-  induction m.
-  - cbn. apply Htau.
-  - cbn. change (fin (S m + n)) with (fin (S (m + n))) in x.
-    destruct x as [x|].
-    + cbn. apply IHm.
-      intros ?. apply Hsigma.
-    + cbn. apply Hsigma.
-Qed.
-
-Definition zero_p {m : nat} {n} : fin m -> fin (m + n).
-Proof.
-  induction m.
-  - intros [].
-  - intros [x|].
-    + exact (shift_p 1 (IHm x)).
-    + exact var_zero.
-Defined.
-
-Lemma scons_p_head' {X} {m n} (f : fin m -> X) (g : fin n -> X) z:
-  (scons_p  f g) (zero_p  z) = f z.
-Proof.
- induction m.
-  - inversion z.
-  - destruct z.
-    + simpl. simpl. now rewrite IHm.
-    + reflexivity.
-Qed.
-
-Lemma scons_p_head_pointwise {X} {m n} (f : fin m -> X) (g : fin n -> X) :
-  pointwise_relation _ eq (funcomp (scons_p f g) zero_p) f.
-Proof.
-  intros z.
-  unfold funcomp.
-  induction m.
-  - inversion z.
-  - destruct z.
-    + simpl. now rewrite IHm.
-    + reflexivity.
-Qed.
-
-Lemma scons_p_tail' X  m n (f : fin m -> X) (g : fin n -> X) z :
-  scons_p  f g (shift_p m z) = g z.
-Proof. induction m; cbn; eauto. Qed.
-
-Lemma scons_p_tail_pointwise X  m n (f : fin m -> X) (g : fin n -> X) :
-  pointwise_relation _ eq (funcomp (scons_p f g) (shift_p m)) g.
-Proof. intros z. induction m; cbn; eauto. Qed.
-
-Lemma destruct_fin {m n} (x : fin (m + n)):
-  (exists x', x = zero_p  x') \/ exists x', x = shift_p m x'.
-Proof.
-  induction m; simpl in *.
-  - right. eauto.
-  - destruct x as [x|].
-    + destruct (IHm x) as [[x' ->] |[x' ->]].
-      * left. now exists (Some x').
-      * right. eauto.
-    + left. exists None. eauto.
-Qed.
-
-Lemma scons_p_comp' X Y m n (f : fin m -> X) (g : fin n -> X) (h : X -> Y) :
-  pointwise_relation _ eq (funcomp h (scons_p f g)) (scons_p (f >> h) (g >> h)).
-Proof.
-  intros x.
-  destruct (destruct_fin x) as [[x' ->]|[x' ->]].
-  (* TODO better way to solve this? *)
-  - revert x'.
-    apply pointwise_forall.
-    change (fun x => (scons_p f g >> h) (zero_p x)) with (zero_p >> scons_p f g >> h).
-    now setoid_rewrite scons_p_head_pointwise.
-  - revert x'.
-    apply pointwise_forall.
-    change (fun x => (scons_p f g >> h) (shift_p m x)) with (shift_p m >> scons_p f g >> h).
-    change (fun x => scons_p (f >> h) (g >> h) (shift_p m x)) with (shift_p m >> scons_p (f >> h) (g >> h)).
-    now rewrite !scons_p_tail_pointwise.
-Qed.
-
-
-Lemma scons_p_congr {X} {m n} (f f' : fin m -> X) (g g': fin n -> X) z:
-  (forall x, f x = f' x) -> (forall x, g x = g' x) -> scons_p f g z = scons_p f' g' z.
-Proof. intros H1 H2. induction m; eauto. cbn. destruct z; eauto. Qed.
-
-(** Generic n-ary lifting operation. *)
-Definition upRen_p p { m : nat } { n : nat } (xi : (fin) (m) -> (fin) (n)) : fin (p + m) -> fin (p + n)  :=
-   scons_p  (zero_p ) (xi >> shift_p _).
-
-Arguments upRen_p p {m n} xi.
-
-(** Generic proof for composition of n-ary lifting. *)
-Lemma up_ren_ren_p p k l m (xi: ren k l) (zeta : ren l m) (rho: ren k m) (E: forall x, (xi >> zeta) x = rho x) :
-  forall x, (upRen_p p xi >> upRen_p p zeta) x = upRen_p p rho x.
-Proof.
-  intros x. destruct (destruct_fin x) as [[? ->]|[? ->]].
-  - unfold upRen_p. unfold funcomp. now repeat rewrite scons_p_head'.
-  - unfold upRen_p. unfold funcomp. repeat rewrite scons_p_tail'.
-    now rewrite <- E.
-Qed.
-
-
-Arguments zero_p m {n}.
-Arguments scons_p  {X} m {n} f g.
-
-Lemma scons_p_eta {X} {m n} {f : fin m -> X}
-      {g : fin n -> X} (h: fin (m + n) -> X) {z: fin (m + n)}:
-  (forall x, g x = h (shift_p m x)) -> (forall x, f x = h (zero_p m x)) -> scons_p m f g z = h z.
-Proof.
-  intros H1 H2. destruct (destruct_fin z) as [[? ->] |[? ->]].
-  - rewrite scons_p_head'. eauto.
-  - rewrite scons_p_tail'. eauto.
-Qed.
-
-Arguments scons_p_eta {X} {m n} {f g} h {z}.
-Arguments scons_p_congr {X} {m n} {f f'} {g g'} {z}.
-
-(** ** Notations for Scoped Syntax *)
-
-Module ScopedNotations.
-  Include RenNotations.
-  Include SubstNotations.
-  Include CombineNotations.
-
-(* Notation "s , sigma" := (scons s sigma) (at level 60, format "s ,  sigma", right associativity) : subst_scope. *)
-
-  Notation "s '..'" := (scons s ids) (at level 1, format "s ..") : subst_scope.
-
-  Notation "↑" := (shift) : subst_scope.
-
-  #[global]
-  Open Scope fscope.
-  #[global]
-  Open Scope subst_scope.
-End ScopedNotations.
-
-
-(** ** Tactics for Scoped Syntax *)
-
-Tactic Notation "auto_case" tactic(t) :=  (match goal with
-                                           | [|- forall (i : fin 0), _] => intros []; t
-                                           | [|- forall (i : fin (S (S (S (S _))))), _] => intros [[[[|]|]|]|]; t
-                                           | [|- forall (i : fin (S (S (S _)))), _] => intros [[[|]|]|]; t
-                                           | [|- forall (i : fin (S (S _))), _] => intros [[?|]|]; t
-                                           | [|- forall (i : fin (S _)), _] =>  intros [?|]; t
-                                           end).
-
-#[export] Hint Rewrite @scons_p_head' @scons_p_tail' : FunctorInstances.
-
-(** Generic fsimpl tactic: simplifies the above primitives in a goal. *)
-Ltac fsimpl :=
-  repeat match goal with
-         | [|- context[id >> ?f]] => change (id >> f) with f (* AsimplCompIdL *)
-         | [|- context[?f >> id]] => change (f >> id) with f (* AsimplCompIdR *)
-         | [|- context [id ?s]] => change (id s) with s
-         | [|- context[(?f >> ?g) >> ?h]] => change ((f >> g) >> h) with (f >> (g >> h)) (* AsimplComp *)
-         (* | [|- zero_p >> scons_p ?f ?g] => rewrite scons_p_head *)
-         | [|- context[(?s .: ?sigma) var_zero]] => change ((s.:sigma) var_zero) with s
-         | [|- context[(?s .: ?sigma) (shift ?m)]] => change ((s.:sigma) (shift m)) with (sigma m)
-           (* first [progress setoid_rewrite scons_tail' | progress setoid_rewrite scons_tail'_pointwise ] *)
-         | [|- context[idren >> ?f]] => change (idren >> f) with f
-         | [|- context[?f >> idren]] => change (f >> idren) with f
-         | [|- context[?f >> (?x .: ?g)]] => change (f >> (x .: g)) with g (* f should evaluate to shift *)
-         | [|- context[?x2 .: (funcomp ?f shift)]] => change (scons x2 (funcomp f shift)) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite (@scons_eta' _ _ f); eta_reduce
-         | [|- context[?f var_zero .: ?g]] => change (scons (f var_zero) g) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite scons_eta'; eta_reduce
-         | [|- _ =  ?h (?f ?s)] => change (h (f s)) with ((f >> h) s)
-         | [|-  ?h (?f ?s) = _] => change (h (f s)) with ((f >> h) s)
-         | [|- context[funcomp _ (scons _ _)]] => setoid_rewrite scons_comp'; eta_reduce
-         | [|- context[funcomp _ (scons_p _ _ _)]] => setoid_rewrite scons_p_comp'; eta_reduce
-         | [|- context[scons (@var_zero _) shift]] => setoid_rewrite scons_eta_id'; eta_reduce
-         (* | _ => progress autorewrite with FunctorInstances *)
-         | [|- context[funcomp (scons_p _ _ _) (zero_p _)]] =>
-           first [progress setoid_rewrite scons_p_head_pointwise | progress setoid_rewrite scons_p_head' ]
-         | [|- context[scons_p _ _ _ (zero_p _ _)]] => setoid_rewrite scons_p_head'
-         | [|- context[funcomp (scons_p _ _ _) (shift_p _)]] =>
-           first [progress setoid_rewrite scons_p_tail_pointwise | progress setoid_rewrite scons_p_tail' ]
-         | [|- context[scons_p _ _ _ (shift_p _ _)]] => setoid_rewrite scons_p_tail'
-         | _ => first [progress minimize | progress cbn [shift scons scons_p] ]
-         end.
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index 75ef645..659d8b0 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -1,4 +1,4 @@
-Require Import core fintype.
+Require Import core unscoped.
 
 Require Import Setoid Morphisms Relation_Definitions.
 
@@ -33,345 +33,257 @@ Proof.
 exact (eq_refl).
 Qed.
 
-Inductive PTm (n_PTm : nat) : Type :=
-  | VarPTm : fin n_PTm -> PTm n_PTm
-  | PAbs : PTm (S n_PTm) -> PTm n_PTm
-  | PApp : PTm n_PTm -> PTm n_PTm -> PTm n_PTm
-  | PPair : PTm n_PTm -> PTm n_PTm -> PTm n_PTm
-  | PProj : PTag -> PTm n_PTm -> PTm n_PTm
-  | PConst : TTag -> PTm n_PTm
-  | PUniv : nat -> PTm n_PTm
-  | PBot : PTm n_PTm.
+Inductive PTm : Type :=
+  | VarPTm : nat -> PTm
+  | PAbs : PTm -> PTm
+  | PApp : PTm -> PTm -> PTm
+  | PPair : PTm -> PTm -> PTm
+  | PProj : PTag -> PTm -> PTm
+  | PConst : TTag -> PTm
+  | PUniv : nat -> PTm
+  | PBot : PTm.
 
-Lemma congr_PAbs {m_PTm : nat} {s0 : PTm (S m_PTm)} {t0 : PTm (S m_PTm)}
-  (H0 : s0 = t0) : PAbs m_PTm s0 = PAbs m_PTm t0.
+Lemma congr_PAbs {s0 : PTm} {t0 : PTm} (H0 : s0 = t0) : PAbs s0 = PAbs t0.
 Proof.
-exact (eq_trans eq_refl (ap (fun x => PAbs m_PTm x) H0)).
+exact (eq_trans eq_refl (ap (fun x => PAbs x) H0)).
 Qed.
 
-Lemma congr_PApp {m_PTm : nat} {s0 : PTm m_PTm} {s1 : PTm m_PTm}
-  {t0 : PTm m_PTm} {t1 : PTm m_PTm} (H0 : s0 = t0) (H1 : s1 = t1) :
-  PApp m_PTm s0 s1 = PApp m_PTm t0 t1.
+Lemma congr_PApp {s0 : PTm} {s1 : PTm} {t0 : PTm} {t1 : PTm} (H0 : s0 = t0)
+  (H1 : s1 = t1) : PApp s0 s1 = PApp t0 t1.
 Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => PApp m_PTm x s1) H0))
-         (ap (fun x => PApp m_PTm t0 x) H1)).
+exact (eq_trans (eq_trans eq_refl (ap (fun x => PApp x s1) H0))
+         (ap (fun x => PApp t0 x) H1)).
 Qed.
 
-Lemma congr_PPair {m_PTm : nat} {s0 : PTm m_PTm} {s1 : PTm m_PTm}
-  {t0 : PTm m_PTm} {t1 : PTm m_PTm} (H0 : s0 = t0) (H1 : s1 = t1) :
-  PPair m_PTm s0 s1 = PPair m_PTm t0 t1.
+Lemma congr_PPair {s0 : PTm} {s1 : PTm} {t0 : PTm} {t1 : PTm} (H0 : s0 = t0)
+  (H1 : s1 = t1) : PPair s0 s1 = PPair t0 t1.
 Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => PPair m_PTm x s1) H0))
-         (ap (fun x => PPair m_PTm t0 x) H1)).
+exact (eq_trans (eq_trans eq_refl (ap (fun x => PPair x s1) H0))
+         (ap (fun x => PPair t0 x) H1)).
 Qed.
 
-Lemma congr_PProj {m_PTm : nat} {s0 : PTag} {s1 : PTm m_PTm} {t0 : PTag}
-  {t1 : PTm m_PTm} (H0 : s0 = t0) (H1 : s1 = t1) :
-  PProj m_PTm s0 s1 = PProj m_PTm t0 t1.
+Lemma congr_PProj {s0 : PTag} {s1 : PTm} {t0 : PTag} {t1 : PTm}
+  (H0 : s0 = t0) (H1 : s1 = t1) : PProj s0 s1 = PProj t0 t1.
 Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => PProj m_PTm x s1) H0))
-         (ap (fun x => PProj m_PTm t0 x) H1)).
+exact (eq_trans (eq_trans eq_refl (ap (fun x => PProj x s1) H0))
+         (ap (fun x => PProj t0 x) H1)).
 Qed.
 
-Lemma congr_PConst {m_PTm : nat} {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
-  PConst m_PTm s0 = PConst m_PTm t0.
+Lemma congr_PConst {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
+  PConst s0 = PConst t0.
 Proof.
-exact (eq_trans eq_refl (ap (fun x => PConst m_PTm x) H0)).
+exact (eq_trans eq_refl (ap (fun x => PConst x) H0)).
 Qed.
 
-Lemma congr_PUniv {m_PTm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
-  PUniv m_PTm s0 = PUniv m_PTm t0.
+Lemma congr_PUniv {s0 : nat} {t0 : nat} (H0 : s0 = t0) : PUniv s0 = PUniv t0.
 Proof.
-exact (eq_trans eq_refl (ap (fun x => PUniv m_PTm x) H0)).
+exact (eq_trans eq_refl (ap (fun x => PUniv x) H0)).
 Qed.
 
-Lemma congr_PBot {m_PTm : nat} : PBot m_PTm = PBot m_PTm.
+Lemma congr_PBot : PBot = PBot.
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma upRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n) :
-  fin (S m) -> fin (S n).
+Lemma upRen_PTm_PTm (xi : nat -> nat) : nat -> nat.
 Proof.
 exact (up_ren xi).
 Defined.
 
-Lemma upRen_list_PTm_PTm (p : nat) {m : nat} {n : nat} (xi : fin m -> fin n)
-  : fin (plus p m) -> fin (plus p n).
-Proof.
-exact (upRen_p p xi).
-Defined.
-
-Fixpoint ren_PTm {m_PTm : nat} {n_PTm : nat}
-(xi_PTm : fin m_PTm -> fin n_PTm) (s : PTm m_PTm) {struct s} : PTm n_PTm :=
+Fixpoint ren_PTm (xi_PTm : nat -> nat) (s : PTm) {struct s} : PTm :=
   match s with
-  | VarPTm _ s0 => VarPTm n_PTm (xi_PTm s0)
-  | PAbs _ s0 => PAbs n_PTm (ren_PTm (upRen_PTm_PTm xi_PTm) s0)
-  | PApp _ s0 s1 => PApp n_PTm (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
-  | PPair _ s0 s1 => PPair n_PTm (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
-  | PProj _ s0 s1 => PProj n_PTm s0 (ren_PTm xi_PTm s1)
-  | PConst _ s0 => PConst n_PTm s0
-  | PUniv _ s0 => PUniv n_PTm s0
-  | PBot _ => PBot n_PTm
+  | VarPTm s0 => VarPTm (xi_PTm s0)
+  | PAbs s0 => PAbs (ren_PTm (upRen_PTm_PTm xi_PTm) s0)
+  | PApp s0 s1 => PApp (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
+  | PPair s0 s1 => PPair (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
+  | PProj s0 s1 => PProj s0 (ren_PTm xi_PTm s1)
+  | PConst s0 => PConst s0
+  | PUniv s0 => PUniv s0
+  | PBot => PBot
   end.
 
-Lemma up_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm) :
-  fin (S m) -> PTm (S n_PTm).
+Lemma up_PTm_PTm (sigma : nat -> PTm) : nat -> PTm.
 Proof.
-exact (scons (VarPTm (S n_PTm) var_zero) (funcomp (ren_PTm shift) sigma)).
+exact (scons (VarPTm var_zero) (funcomp (ren_PTm shift) sigma)).
 Defined.
 
-Lemma up_list_PTm_PTm (p : nat) {m : nat} {n_PTm : nat}
-  (sigma : fin m -> PTm n_PTm) : fin (plus p m) -> PTm (plus p n_PTm).
-Proof.
-exact (scons_p p (funcomp (VarPTm (plus p n_PTm)) (zero_p p))
-         (funcomp (ren_PTm (shift_p p)) sigma)).
-Defined.
-
-Fixpoint subst_PTm {m_PTm : nat} {n_PTm : nat}
-(sigma_PTm : fin m_PTm -> PTm n_PTm) (s : PTm m_PTm) {struct s} : PTm n_PTm
-:=
+Fixpoint subst_PTm (sigma_PTm : nat -> PTm) (s : PTm) {struct s} : PTm :=
   match s with
-  | VarPTm _ s0 => sigma_PTm s0
-  | PAbs _ s0 => PAbs n_PTm (subst_PTm (up_PTm_PTm sigma_PTm) s0)
-  | PApp _ s0 s1 =>
-      PApp n_PTm (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
-  | PPair _ s0 s1 =>
-      PPair n_PTm (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
-  | PProj _ s0 s1 => PProj n_PTm s0 (subst_PTm sigma_PTm s1)
-  | PConst _ s0 => PConst n_PTm s0
-  | PUniv _ s0 => PUniv n_PTm s0
-  | PBot _ => PBot n_PTm
+  | VarPTm s0 => sigma_PTm s0
+  | PAbs s0 => PAbs (subst_PTm (up_PTm_PTm sigma_PTm) s0)
+  | PApp s0 s1 => PApp (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
+  | PPair s0 s1 => PPair (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
+  | PProj s0 s1 => PProj s0 (subst_PTm sigma_PTm s1)
+  | PConst s0 => PConst s0
+  | PUniv s0 => PUniv s0
+  | PBot => PBot
   end.
 
-Lemma upId_PTm_PTm {m_PTm : nat} (sigma : fin m_PTm -> PTm m_PTm)
-  (Eq : forall x, sigma x = VarPTm m_PTm x) :
-  forall x, up_PTm_PTm sigma x = VarPTm (S m_PTm) x.
+Lemma upId_PTm_PTm (sigma : nat -> PTm) (Eq : forall x, sigma x = VarPTm x) :
+  forall x, up_PTm_PTm sigma x = VarPTm x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_PTm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma upId_list_PTm_PTm {p : nat} {m_PTm : nat}
-  (sigma : fin m_PTm -> PTm m_PTm) (Eq : forall x, sigma x = VarPTm m_PTm x)
-  : forall x, up_list_PTm_PTm p sigma x = VarPTm (plus p m_PTm) x.
-Proof.
-exact (fun n =>
-       scons_p_eta (VarPTm (plus p m_PTm))
-         (fun n => ap (ren_PTm (shift_p p)) (Eq n)) (fun n => eq_refl)).
-Qed.
-
-Fixpoint idSubst_PTm {m_PTm : nat} (sigma_PTm : fin m_PTm -> PTm m_PTm)
-(Eq_PTm : forall x, sigma_PTm x = VarPTm m_PTm x) (s : PTm m_PTm) {struct s}
-   : subst_PTm sigma_PTm s = s :=
+Fixpoint idSubst_PTm (sigma_PTm : nat -> PTm)
+(Eq_PTm : forall x, sigma_PTm x = VarPTm x) (s : PTm) {struct s} :
+subst_PTm sigma_PTm s = s :=
   match s with
-  | VarPTm _ s0 => Eq_PTm s0
-  | PAbs _ s0 =>
+  | VarPTm s0 => Eq_PTm s0
+  | PAbs s0 =>
       congr_PAbs
         (idSubst_PTm (up_PTm_PTm sigma_PTm) (upId_PTm_PTm _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (idSubst_PTm sigma_PTm Eq_PTm s0)
         (idSubst_PTm sigma_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (idSubst_PTm sigma_PTm Eq_PTm s0)
         (idSubst_PTm sigma_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
-      congr_PProj (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PProj s0 s1 => congr_PProj (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1)
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma upExtRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n)
-  (zeta : fin m -> fin n) (Eq : forall x, xi x = zeta x) :
+Lemma upExtRen_PTm_PTm (xi : nat -> nat) (zeta : nat -> nat)
+  (Eq : forall x, xi x = zeta x) :
   forall x, upRen_PTm_PTm xi x = upRen_PTm_PTm zeta x.
 Proof.
-exact (fun n =>
-       match n with
-       | Some fin_n => ap shift (Eq fin_n)
-       | None => eq_refl
-       end).
+exact (fun n => match n with
+                | S n' => ap shift (Eq n')
+                | O => eq_refl
+                end).
 Qed.
 
-Lemma upExtRen_list_PTm_PTm {p : nat} {m : nat} {n : nat}
-  (xi : fin m -> fin n) (zeta : fin m -> fin n)
-  (Eq : forall x, xi x = zeta x) :
-  forall x, upRen_list_PTm_PTm p xi x = upRen_list_PTm_PTm p zeta x.
-Proof.
-exact (fun n =>
-       scons_p_congr (fun n => eq_refl) (fun n => ap (shift_p p) (Eq n))).
-Qed.
-
-Fixpoint extRen_PTm {m_PTm : nat} {n_PTm : nat}
-(xi_PTm : fin m_PTm -> fin n_PTm) (zeta_PTm : fin m_PTm -> fin n_PTm)
-(Eq_PTm : forall x, xi_PTm x = zeta_PTm x) (s : PTm m_PTm) {struct s} :
+Fixpoint extRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat)
+(Eq_PTm : forall x, xi_PTm x = zeta_PTm x) (s : PTm) {struct s} :
 ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
   match s with
-  | VarPTm _ s0 => ap (VarPTm n_PTm) (Eq_PTm s0)
-  | PAbs _ s0 =>
+  | VarPTm s0 => ap (VarPTm) (Eq_PTm s0)
+  | PAbs s0 =>
       congr_PAbs
         (extRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
            (upExtRen_PTm_PTm _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0)
         (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0)
         (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0) (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma upExt_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm)
-  (tau : fin m -> PTm n_PTm) (Eq : forall x, sigma x = tau x) :
+Lemma upExt_PTm_PTm (sigma : nat -> PTm) (tau : nat -> PTm)
+  (Eq : forall x, sigma x = tau x) :
   forall x, up_PTm_PTm sigma x = up_PTm_PTm tau x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_PTm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma upExt_list_PTm_PTm {p : nat} {m : nat} {n_PTm : nat}
-  (sigma : fin m -> PTm n_PTm) (tau : fin m -> PTm n_PTm)
-  (Eq : forall x, sigma x = tau x) :
-  forall x, up_list_PTm_PTm p sigma x = up_list_PTm_PTm p tau x.
-Proof.
-exact (fun n =>
-       scons_p_congr (fun n => eq_refl)
-         (fun n => ap (ren_PTm (shift_p p)) (Eq n))).
-Qed.
-
-Fixpoint ext_PTm {m_PTm : nat} {n_PTm : nat}
-(sigma_PTm : fin m_PTm -> PTm n_PTm) (tau_PTm : fin m_PTm -> PTm n_PTm)
-(Eq_PTm : forall x, sigma_PTm x = tau_PTm x) (s : PTm m_PTm) {struct s} :
+Fixpoint ext_PTm (sigma_PTm : nat -> PTm) (tau_PTm : nat -> PTm)
+(Eq_PTm : forall x, sigma_PTm x = tau_PTm x) (s : PTm) {struct s} :
 subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
   match s with
-  | VarPTm _ s0 => Eq_PTm s0
-  | PAbs _ s0 =>
+  | VarPTm s0 => Eq_PTm s0
+  | PAbs s0 =>
       congr_PAbs
         (ext_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
            (upExt_PTm_PTm _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (ext_PTm sigma_PTm tau_PTm Eq_PTm s0)
         (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (ext_PTm sigma_PTm tau_PTm Eq_PTm s0)
         (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0) (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma up_ren_ren_PTm_PTm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
-  (zeta : fin l -> fin m) (rho : fin k -> fin m)
-  (Eq : forall x, funcomp zeta xi x = rho x) :
+Lemma up_ren_ren_PTm_PTm (xi : nat -> nat) (zeta : nat -> nat)
+  (rho : nat -> nat) (Eq : forall x, funcomp zeta xi x = rho x) :
   forall x,
   funcomp (upRen_PTm_PTm zeta) (upRen_PTm_PTm xi) x = upRen_PTm_PTm rho x.
 Proof.
 exact (up_ren_ren xi zeta rho Eq).
 Qed.
 
-Lemma up_ren_ren_list_PTm_PTm {p : nat} {k : nat} {l : nat} {m : nat}
-  (xi : fin k -> fin l) (zeta : fin l -> fin m) (rho : fin k -> fin m)
-  (Eq : forall x, funcomp zeta xi x = rho x) :
-  forall x,
-  funcomp (upRen_list_PTm_PTm p zeta) (upRen_list_PTm_PTm p xi) x =
-  upRen_list_PTm_PTm p rho x.
-Proof.
-exact (up_ren_ren_p Eq).
-Qed.
-
-Fixpoint compRenRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-(xi_PTm : fin m_PTm -> fin k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
-(rho_PTm : fin m_PTm -> fin l_PTm)
-(Eq_PTm : forall x, funcomp zeta_PTm xi_PTm x = rho_PTm x) (s : PTm m_PTm)
-{struct s} : ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm rho_PTm s :=
+Fixpoint compRenRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat)
+(rho_PTm : nat -> nat)
+(Eq_PTm : forall x, funcomp zeta_PTm xi_PTm x = rho_PTm x) (s : PTm) {struct
+ s} : ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm rho_PTm s :=
   match s with
-  | VarPTm _ s0 => ap (VarPTm l_PTm) (Eq_PTm s0)
-  | PAbs _ s0 =>
+  | VarPTm s0 => ap (VarPTm) (Eq_PTm s0)
+  | PAbs s0 =>
       congr_PAbs
         (compRenRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
            (upRen_PTm_PTm rho_PTm) (up_ren_ren _ _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0)
         (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0)
         (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0)
         (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma up_ren_subst_PTm_PTm {k : nat} {l : nat} {m_PTm : nat}
-  (xi : fin k -> fin l) (tau : fin l -> PTm m_PTm)
-  (theta : fin k -> PTm m_PTm) (Eq : forall x, funcomp tau xi x = theta x) :
+Lemma up_ren_subst_PTm_PTm (xi : nat -> nat) (tau : nat -> PTm)
+  (theta : nat -> PTm) (Eq : forall x, funcomp tau xi x = theta x) :
   forall x,
   funcomp (up_PTm_PTm tau) (upRen_PTm_PTm xi) x = up_PTm_PTm theta x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_PTm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma up_ren_subst_list_PTm_PTm {p : nat} {k : nat} {l : nat} {m_PTm : nat}
-  (xi : fin k -> fin l) (tau : fin l -> PTm m_PTm)
-  (theta : fin k -> PTm m_PTm) (Eq : forall x, funcomp tau xi x = theta x) :
-  forall x,
-  funcomp (up_list_PTm_PTm p tau) (upRen_list_PTm_PTm p xi) x =
-  up_list_PTm_PTm p theta x.
-Proof.
-exact (fun n =>
-       eq_trans (scons_p_comp' _ _ _ n)
-         (scons_p_congr (fun z => scons_p_head' _ _ z)
-            (fun z =>
-             eq_trans (scons_p_tail' _ _ (xi z))
-               (ap (ren_PTm (shift_p p)) (Eq z))))).
-Qed.
-
-Fixpoint compRenSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-(xi_PTm : fin m_PTm -> fin k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
-(theta_PTm : fin m_PTm -> PTm l_PTm)
-(Eq_PTm : forall x, funcomp tau_PTm xi_PTm x = theta_PTm x) (s : PTm m_PTm)
-{struct s} : subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm theta_PTm s :=
+Fixpoint compRenSubst_PTm (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm)
+(theta_PTm : nat -> PTm)
+(Eq_PTm : forall x, funcomp tau_PTm xi_PTm x = theta_PTm x) (s : PTm) {struct
+ s} : subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm theta_PTm s :=
   match s with
-  | VarPTm _ s0 => Eq_PTm s0
-  | PAbs _ s0 =>
+  | VarPTm s0 => Eq_PTm s0
+  | PAbs s0 =>
       congr_PAbs
         (compRenSubst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm tau_PTm)
            (up_PTm_PTm theta_PTm) (up_ren_subst_PTm_PTm _ _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0)
         (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0)
         (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0)
         (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma up_subst_ren_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
-  (sigma : fin k -> PTm l_PTm) (zeta_PTm : fin l_PTm -> fin m_PTm)
-  (theta : fin k -> PTm m_PTm)
+Lemma up_subst_ren_PTm_PTm (sigma : nat -> PTm) (zeta_PTm : nat -> nat)
+  (theta : nat -> PTm)
   (Eq : forall x, funcomp (ren_PTm zeta_PTm) sigma x = theta x) :
   forall x,
   funcomp (ren_PTm (upRen_PTm_PTm zeta_PTm)) (up_PTm_PTm sigma) x =
@@ -379,72 +291,46 @@ Lemma up_subst_ren_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n =>
+       | S n' =>
            eq_trans
              (compRenRen_PTm shift (upRen_PTm_PTm zeta_PTm)
-                (funcomp shift zeta_PTm) (fun x => eq_refl) (sigma fin_n))
+                (funcomp shift zeta_PTm) (fun x => eq_refl) (sigma n'))
              (eq_trans
                 (eq_sym
                    (compRenRen_PTm zeta_PTm shift (funcomp shift zeta_PTm)
-                      (fun x => eq_refl) (sigma fin_n)))
-                (ap (ren_PTm shift) (Eq fin_n)))
-       | None => eq_refl
+                      (fun x => eq_refl) (sigma n')))
+                (ap (ren_PTm shift) (Eq n')))
+       | O => eq_refl
        end).
 Qed.
 
-Lemma up_subst_ren_list_PTm_PTm {p : nat} {k : nat} {l_PTm : nat}
-  {m_PTm : nat} (sigma : fin k -> PTm l_PTm)
-  (zeta_PTm : fin l_PTm -> fin m_PTm) (theta : fin k -> PTm m_PTm)
-  (Eq : forall x, funcomp (ren_PTm zeta_PTm) sigma x = theta x) :
-  forall x,
-  funcomp (ren_PTm (upRen_list_PTm_PTm p zeta_PTm)) (up_list_PTm_PTm p sigma)
-    x = up_list_PTm_PTm p theta x.
-Proof.
-exact (fun n =>
-       eq_trans (scons_p_comp' _ _ _ n)
-         (scons_p_congr
-            (fun x => ap (VarPTm (plus p m_PTm)) (scons_p_head' _ _ x))
-            (fun n =>
-             eq_trans
-               (compRenRen_PTm (shift_p p) (upRen_list_PTm_PTm p zeta_PTm)
-                  (funcomp (shift_p p) zeta_PTm)
-                  (fun x => scons_p_tail' _ _ x) (sigma n))
-               (eq_trans
-                  (eq_sym
-                     (compRenRen_PTm zeta_PTm (shift_p p)
-                        (funcomp (shift_p p) zeta_PTm) (fun x => eq_refl)
-                        (sigma n))) (ap (ren_PTm (shift_p p)) (Eq n)))))).
-Qed.
-
-Fixpoint compSubstRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-(sigma_PTm : fin m_PTm -> PTm k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
-(theta_PTm : fin m_PTm -> PTm l_PTm)
+Fixpoint compSubstRen_PTm (sigma_PTm : nat -> PTm) (zeta_PTm : nat -> nat)
+(theta_PTm : nat -> PTm)
 (Eq_PTm : forall x, funcomp (ren_PTm zeta_PTm) sigma_PTm x = theta_PTm x)
-(s : PTm m_PTm) {struct s} :
+(s : PTm) {struct s} :
 ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
   match s with
-  | VarPTm _ s0 => Eq_PTm s0
-  | PAbs _ s0 =>
+  | VarPTm s0 => Eq_PTm s0
+  | PAbs s0 =>
       congr_PAbs
         (compSubstRen_PTm (up_PTm_PTm sigma_PTm) (upRen_PTm_PTm zeta_PTm)
            (up_PTm_PTm theta_PTm) (up_subst_ren_PTm_PTm _ _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0)
         (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0)
         (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0)
         (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma up_subst_subst_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
-  (sigma : fin k -> PTm l_PTm) (tau_PTm : fin l_PTm -> PTm m_PTm)
-  (theta : fin k -> PTm m_PTm)
+Lemma up_subst_subst_PTm_PTm (sigma : nat -> PTm) (tau_PTm : nat -> PTm)
+  (theta : nat -> PTm)
   (Eq : forall x, funcomp (subst_PTm tau_PTm) sigma x = theta x) :
   forall x,
   funcomp (subst_PTm (up_PTm_PTm tau_PTm)) (up_PTm_PTm sigma) x =
@@ -452,649 +338,501 @@ Lemma up_subst_subst_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n =>
+       | S n' =>
            eq_trans
              (compRenSubst_PTm shift (up_PTm_PTm tau_PTm)
                 (funcomp (up_PTm_PTm tau_PTm) shift) (fun x => eq_refl)
-                (sigma fin_n))
+                (sigma n'))
              (eq_trans
                 (eq_sym
                    (compSubstRen_PTm tau_PTm shift
                       (funcomp (ren_PTm shift) tau_PTm) (fun x => eq_refl)
-                      (sigma fin_n))) (ap (ren_PTm shift) (Eq fin_n)))
-       | None => eq_refl
+                      (sigma n'))) (ap (ren_PTm shift) (Eq n')))
+       | O => eq_refl
        end).
 Qed.
 
-Lemma up_subst_subst_list_PTm_PTm {p : nat} {k : nat} {l_PTm : nat}
-  {m_PTm : nat} (sigma : fin k -> PTm l_PTm)
-  (tau_PTm : fin l_PTm -> PTm m_PTm) (theta : fin k -> PTm m_PTm)
-  (Eq : forall x, funcomp (subst_PTm tau_PTm) sigma x = theta x) :
-  forall x,
-  funcomp (subst_PTm (up_list_PTm_PTm p tau_PTm)) (up_list_PTm_PTm p sigma) x =
-  up_list_PTm_PTm p theta x.
-Proof.
-exact (fun n =>
-       eq_trans
-         (scons_p_comp' (funcomp (VarPTm (plus p l_PTm)) (zero_p p)) _ _ n)
-         (scons_p_congr
-            (fun x => scons_p_head' _ (fun z => ren_PTm (shift_p p) _) x)
-            (fun n =>
-             eq_trans
-               (compRenSubst_PTm (shift_p p) (up_list_PTm_PTm p tau_PTm)
-                  (funcomp (up_list_PTm_PTm p tau_PTm) (shift_p p))
-                  (fun x => eq_refl) (sigma n))
-               (eq_trans
-                  (eq_sym
-                     (compSubstRen_PTm tau_PTm (shift_p p) _
-                        (fun x => eq_sym (scons_p_tail' _ _ x)) (sigma n)))
-                  (ap (ren_PTm (shift_p p)) (Eq n)))))).
-Qed.
-
-Fixpoint compSubstSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-(sigma_PTm : fin m_PTm -> PTm k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
-(theta_PTm : fin m_PTm -> PTm l_PTm)
+Fixpoint compSubstSubst_PTm (sigma_PTm : nat -> PTm) (tau_PTm : nat -> PTm)
+(theta_PTm : nat -> PTm)
 (Eq_PTm : forall x, funcomp (subst_PTm tau_PTm) sigma_PTm x = theta_PTm x)
-(s : PTm m_PTm) {struct s} :
+(s : PTm) {struct s} :
 subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
   match s with
-  | VarPTm _ s0 => Eq_PTm s0
-  | PAbs _ s0 =>
+  | VarPTm s0 => Eq_PTm s0
+  | PAbs s0 =>
       congr_PAbs
         (compSubstSubst_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
            (up_PTm_PTm theta_PTm) (up_subst_subst_PTm_PTm _ _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0)
         (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0)
         (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0)
         (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma renRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
-  (s : PTm m_PTm) :
+Lemma renRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) (s : PTm) :
   ren_PTm zeta_PTm (ren_PTm xi_PTm s) = ren_PTm (funcomp zeta_PTm xi_PTm) s.
 Proof.
 exact (compRenRen_PTm xi_PTm zeta_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma renRen'_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm) :
+Lemma renRen'_PTm_pointwise (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) :
   pointwise_relation _ eq (funcomp (ren_PTm zeta_PTm) (ren_PTm xi_PTm))
     (ren_PTm (funcomp zeta_PTm xi_PTm)).
 Proof.
 exact (fun s => compRenRen_PTm xi_PTm zeta_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma renSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
-  (s : PTm m_PTm) :
+Lemma renSubst_PTm (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm) (s : PTm) :
   subst_PTm tau_PTm (ren_PTm xi_PTm s) = subst_PTm (funcomp tau_PTm xi_PTm) s.
 Proof.
 exact (compRenSubst_PTm xi_PTm tau_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma renSubst_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm) :
+Lemma renSubst_PTm_pointwise (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm) :
   pointwise_relation _ eq (funcomp (subst_PTm tau_PTm) (ren_PTm xi_PTm))
     (subst_PTm (funcomp tau_PTm xi_PTm)).
 Proof.
 exact (fun s => compRenSubst_PTm xi_PTm tau_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (sigma_PTm : fin m_PTm -> PTm k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm)
-  (s : PTm m_PTm) :
+Lemma substRen_PTm (sigma_PTm : nat -> PTm) (zeta_PTm : nat -> nat) (s : PTm)
+  :
   ren_PTm zeta_PTm (subst_PTm sigma_PTm s) =
   subst_PTm (funcomp (ren_PTm zeta_PTm) sigma_PTm) s.
 Proof.
 exact (compSubstRen_PTm sigma_PTm zeta_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substRen_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (sigma_PTm : fin m_PTm -> PTm k_PTm) (zeta_PTm : fin k_PTm -> fin l_PTm) :
+Lemma substRen_PTm_pointwise (sigma_PTm : nat -> PTm) (zeta_PTm : nat -> nat)
+  :
   pointwise_relation _ eq (funcomp (ren_PTm zeta_PTm) (subst_PTm sigma_PTm))
     (subst_PTm (funcomp (ren_PTm zeta_PTm) sigma_PTm)).
 Proof.
 exact (fun s => compSubstRen_PTm sigma_PTm zeta_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (sigma_PTm : fin m_PTm -> PTm k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm)
-  (s : PTm m_PTm) :
+Lemma substSubst_PTm (sigma_PTm : nat -> PTm) (tau_PTm : nat -> PTm)
+  (s : PTm) :
   subst_PTm tau_PTm (subst_PTm sigma_PTm s) =
   subst_PTm (funcomp (subst_PTm tau_PTm) sigma_PTm) s.
 Proof.
 exact (compSubstSubst_PTm sigma_PTm tau_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substSubst_PTm_pointwise {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
-  (sigma_PTm : fin m_PTm -> PTm k_PTm) (tau_PTm : fin k_PTm -> PTm l_PTm) :
+Lemma substSubst_PTm_pointwise (sigma_PTm : nat -> PTm)
+  (tau_PTm : nat -> PTm) :
   pointwise_relation _ eq (funcomp (subst_PTm tau_PTm) (subst_PTm sigma_PTm))
     (subst_PTm (funcomp (subst_PTm tau_PTm) sigma_PTm)).
 Proof.
 exact (fun s => compSubstSubst_PTm sigma_PTm tau_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma rinstInst_up_PTm_PTm {m : nat} {n_PTm : nat} (xi : fin m -> fin n_PTm)
-  (sigma : fin m -> PTm n_PTm)
-  (Eq : forall x, funcomp (VarPTm n_PTm) xi x = sigma x) :
-  forall x,
-  funcomp (VarPTm (S n_PTm)) (upRen_PTm_PTm xi) x = up_PTm_PTm sigma x.
+Lemma rinstInst_up_PTm_PTm (xi : nat -> nat) (sigma : nat -> PTm)
+  (Eq : forall x, funcomp (VarPTm) xi x = sigma x) :
+  forall x, funcomp (VarPTm) (upRen_PTm_PTm xi) x = up_PTm_PTm sigma x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_PTm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_PTm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma rinstInst_up_list_PTm_PTm {p : nat} {m : nat} {n_PTm : nat}
-  (xi : fin m -> fin n_PTm) (sigma : fin m -> PTm n_PTm)
-  (Eq : forall x, funcomp (VarPTm n_PTm) xi x = sigma x) :
-  forall x,
-  funcomp (VarPTm (plus p n_PTm)) (upRen_list_PTm_PTm p xi) x =
-  up_list_PTm_PTm p sigma x.
-Proof.
-exact (fun n =>
-       eq_trans (scons_p_comp' _ _ (VarPTm (plus p n_PTm)) n)
-         (scons_p_congr (fun z => eq_refl)
-            (fun n => ap (ren_PTm (shift_p p)) (Eq n)))).
-Qed.
-
-Fixpoint rinst_inst_PTm {m_PTm : nat} {n_PTm : nat}
-(xi_PTm : fin m_PTm -> fin n_PTm) (sigma_PTm : fin m_PTm -> PTm n_PTm)
-(Eq_PTm : forall x, funcomp (VarPTm n_PTm) xi_PTm x = sigma_PTm x)
-(s : PTm m_PTm) {struct s} : ren_PTm xi_PTm s = subst_PTm sigma_PTm s :=
+Fixpoint rinst_inst_PTm (xi_PTm : nat -> nat) (sigma_PTm : nat -> PTm)
+(Eq_PTm : forall x, funcomp (VarPTm) xi_PTm x = sigma_PTm x) (s : PTm)
+{struct s} : ren_PTm xi_PTm s = subst_PTm sigma_PTm s :=
   match s with
-  | VarPTm _ s0 => Eq_PTm s0
-  | PAbs _ s0 =>
+  | VarPTm s0 => Eq_PTm s0
+  | PAbs s0 =>
       congr_PAbs
         (rinst_inst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm sigma_PTm)
            (rinstInst_up_PTm_PTm _ _ Eq_PTm) s0)
-  | PApp _ s0 s1 =>
+  | PApp s0 s1 =>
       congr_PApp (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0)
         (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
-  | PPair _ s0 s1 =>
+  | PPair s0 s1 =>
       congr_PPair (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0)
         (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
-  | PProj _ s0 s1 =>
+  | PProj s0 s1 =>
       congr_PProj (eq_refl s0) (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
-  | PConst _ s0 => congr_PConst (eq_refl s0)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
+  | PConst s0 => congr_PConst (eq_refl s0)
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PBot => congr_PBot
   end.
 
-Lemma rinstInst'_PTm {m_PTm : nat} {n_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin n_PTm) (s : PTm m_PTm) :
-  ren_PTm xi_PTm s = subst_PTm (funcomp (VarPTm n_PTm) xi_PTm) s.
+Lemma rinstInst'_PTm (xi_PTm : nat -> nat) (s : PTm) :
+  ren_PTm xi_PTm s = subst_PTm (funcomp (VarPTm) xi_PTm) s.
 Proof.
 exact (rinst_inst_PTm xi_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma rinstInst'_PTm_pointwise {m_PTm : nat} {n_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin n_PTm) :
+Lemma rinstInst'_PTm_pointwise (xi_PTm : nat -> nat) :
   pointwise_relation _ eq (ren_PTm xi_PTm)
-    (subst_PTm (funcomp (VarPTm n_PTm) xi_PTm)).
+    (subst_PTm (funcomp (VarPTm) xi_PTm)).
 Proof.
 exact (fun s => rinst_inst_PTm xi_PTm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma instId'_PTm {m_PTm : nat} (s : PTm m_PTm) :
-  subst_PTm (VarPTm m_PTm) s = s.
+Lemma instId'_PTm (s : PTm) : subst_PTm (VarPTm) s = s.
 Proof.
-exact (idSubst_PTm (VarPTm m_PTm) (fun n => eq_refl) s).
+exact (idSubst_PTm (VarPTm) (fun n => eq_refl) s).
 Qed.
 
-Lemma instId'_PTm_pointwise {m_PTm : nat} :
-  pointwise_relation _ eq (subst_PTm (VarPTm m_PTm)) id.
+Lemma instId'_PTm_pointwise : pointwise_relation _ eq (subst_PTm (VarPTm)) id.
 Proof.
-exact (fun s => idSubst_PTm (VarPTm m_PTm) (fun n => eq_refl) s).
+exact (fun s => idSubst_PTm (VarPTm) (fun n => eq_refl) s).
 Qed.
 
-Lemma rinstId'_PTm {m_PTm : nat} (s : PTm m_PTm) : ren_PTm id s = s.
+Lemma rinstId'_PTm (s : PTm) : ren_PTm id s = s.
 Proof.
 exact (eq_ind_r (fun t => t = s) (instId'_PTm s) (rinstInst'_PTm id s)).
 Qed.
 
-Lemma rinstId'_PTm_pointwise {m_PTm : nat} :
-  pointwise_relation _ eq (@ren_PTm m_PTm m_PTm id) id.
+Lemma rinstId'_PTm_pointwise : pointwise_relation _ eq (@ren_PTm id) id.
 Proof.
 exact (fun s =>
        eq_ind_r (fun t => t = s) (instId'_PTm s) (rinstInst'_PTm id s)).
 Qed.
 
-Lemma varL'_PTm {m_PTm : nat} {n_PTm : nat}
-  (sigma_PTm : fin m_PTm -> PTm n_PTm) (x : fin m_PTm) :
-  subst_PTm sigma_PTm (VarPTm m_PTm x) = sigma_PTm x.
+Lemma varL'_PTm (sigma_PTm : nat -> PTm) (x : nat) :
+  subst_PTm sigma_PTm (VarPTm x) = sigma_PTm x.
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma varL'_PTm_pointwise {m_PTm : nat} {n_PTm : nat}
-  (sigma_PTm : fin m_PTm -> PTm n_PTm) :
-  pointwise_relation _ eq (funcomp (subst_PTm sigma_PTm) (VarPTm m_PTm))
-    sigma_PTm.
+Lemma varL'_PTm_pointwise (sigma_PTm : nat -> PTm) :
+  pointwise_relation _ eq (funcomp (subst_PTm sigma_PTm) (VarPTm)) sigma_PTm.
 Proof.
 exact (fun x => eq_refl).
 Qed.
 
-Lemma varLRen'_PTm {m_PTm : nat} {n_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin n_PTm) (x : fin m_PTm) :
-  ren_PTm xi_PTm (VarPTm m_PTm x) = VarPTm n_PTm (xi_PTm x).
+Lemma varLRen'_PTm (xi_PTm : nat -> nat) (x : nat) :
+  ren_PTm xi_PTm (VarPTm x) = VarPTm (xi_PTm x).
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma varLRen'_PTm_pointwise {m_PTm : nat} {n_PTm : nat}
-  (xi_PTm : fin m_PTm -> fin n_PTm) :
-  pointwise_relation _ eq (funcomp (ren_PTm xi_PTm) (VarPTm m_PTm))
-    (funcomp (VarPTm n_PTm) xi_PTm).
+Lemma varLRen'_PTm_pointwise (xi_PTm : nat -> nat) :
+  pointwise_relation _ eq (funcomp (ren_PTm xi_PTm) (VarPTm))
+    (funcomp (VarPTm) xi_PTm).
 Proof.
 exact (fun x => eq_refl).
 Qed.
 
-Inductive Tm (n_Tm : nat) : Type :=
-  | VarTm : fin n_Tm -> Tm n_Tm
-  | Abs : Tm (S n_Tm) -> Tm n_Tm
-  | App : Tm n_Tm -> Tm n_Tm -> Tm n_Tm
-  | Pair : Tm n_Tm -> Tm n_Tm -> Tm n_Tm
-  | Proj : PTag -> Tm n_Tm -> Tm n_Tm
-  | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
-  | Univ : nat -> Tm n_Tm
-  | BVal : bool -> Tm n_Tm
-  | Bool : Tm n_Tm
-  | If : Tm n_Tm -> Tm n_Tm -> Tm n_Tm -> Tm n_Tm.
+Inductive Tm : Type :=
+  | VarTm : nat -> Tm
+  | Abs : Tm -> Tm
+  | App : Tm -> Tm -> Tm
+  | Pair : Tm -> Tm -> Tm
+  | Proj : PTag -> Tm -> Tm
+  | TBind : TTag -> Tm -> Tm -> Tm
+  | Univ : nat -> Tm
+  | BVal : bool -> Tm
+  | Bool : Tm
+  | If : Tm -> Tm -> Tm -> Tm.
 
-Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)}
-  (H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0.
+Lemma congr_Abs {s0 : Tm} {t0 : Tm} (H0 : s0 = t0) : Abs s0 = Abs t0.
 Proof.
-exact (eq_trans eq_refl (ap (fun x => Abs m_Tm x) H0)).
+exact (eq_trans eq_refl (ap (fun x => Abs x) H0)).
 Qed.
 
-Lemma congr_App {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {t0 : Tm m_Tm}
-  {t1 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1) :
-  App m_Tm s0 s1 = App m_Tm t0 t1.
+Lemma congr_App {s0 : Tm} {s1 : Tm} {t0 : Tm} {t1 : Tm} (H0 : s0 = t0)
+  (H1 : s1 = t1) : App s0 s1 = App t0 t1.
 Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => App m_Tm x s1) H0))
-         (ap (fun x => App m_Tm t0 x) H1)).
+exact (eq_trans (eq_trans eq_refl (ap (fun x => App x s1) H0))
+         (ap (fun x => App t0 x) H1)).
 Qed.
 
-Lemma congr_Pair {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {t0 : Tm m_Tm}
-  {t1 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1) :
-  Pair m_Tm s0 s1 = Pair m_Tm t0 t1.
+Lemma congr_Pair {s0 : Tm} {s1 : Tm} {t0 : Tm} {t1 : Tm} (H0 : s0 = t0)
+  (H1 : s1 = t1) : Pair s0 s1 = Pair t0 t1.
 Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => Pair m_Tm x s1) H0))
-         (ap (fun x => Pair m_Tm t0 x) H1)).
+exact (eq_trans (eq_trans eq_refl (ap (fun x => Pair x s1) H0))
+         (ap (fun x => Pair t0 x) H1)).
 Qed.
 
-Lemma congr_Proj {m_Tm : nat} {s0 : PTag} {s1 : Tm m_Tm} {t0 : PTag}
-  {t1 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1) :
-  Proj m_Tm s0 s1 = Proj m_Tm t0 t1.
+Lemma congr_Proj {s0 : PTag} {s1 : Tm} {t0 : PTag} {t1 : Tm} (H0 : s0 = t0)
+  (H1 : s1 = t1) : Proj s0 s1 = Proj t0 t1.
 Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => Proj m_Tm x s1) H0))
-         (ap (fun x => Proj m_Tm t0 x) H1)).
+exact (eq_trans (eq_trans eq_refl (ap (fun x => Proj x s1) H0))
+         (ap (fun x => Proj t0 x) H1)).
 Qed.
 
-Lemma congr_TBind {m_Tm : nat} {s0 : TTag} {s1 : Tm m_Tm} {s2 : Tm (S m_Tm)}
-  {t0 : TTag} {t1 : Tm m_Tm} {t2 : Tm (S m_Tm)} (H0 : s0 = t0) (H1 : s1 = t1)
-  (H2 : s2 = t2) : TBind m_Tm s0 s1 s2 = TBind m_Tm t0 t1 t2.
+Lemma congr_TBind {s0 : TTag} {s1 : Tm} {s2 : Tm} {t0 : TTag} {t1 : Tm}
+  {t2 : Tm} (H0 : s0 = t0) (H1 : s1 = t1) (H2 : s2 = t2) :
+  TBind s0 s1 s2 = TBind t0 t1 t2.
 Proof.
 exact (eq_trans
-         (eq_trans (eq_trans eq_refl (ap (fun x => TBind m_Tm x s1 s2) H0))
-            (ap (fun x => TBind m_Tm t0 x s2) H1))
-         (ap (fun x => TBind m_Tm t0 t1 x) H2)).
+         (eq_trans (eq_trans eq_refl (ap (fun x => TBind x s1 s2) H0))
+            (ap (fun x => TBind t0 x s2) H1))
+         (ap (fun x => TBind t0 t1 x) H2)).
 Qed.
 
-Lemma congr_Univ {m_Tm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
-  Univ m_Tm s0 = Univ m_Tm t0.
+Lemma congr_Univ {s0 : nat} {t0 : nat} (H0 : s0 = t0) : Univ s0 = Univ t0.
 Proof.
-exact (eq_trans eq_refl (ap (fun x => Univ m_Tm x) H0)).
+exact (eq_trans eq_refl (ap (fun x => Univ x) H0)).
 Qed.
 
-Lemma congr_BVal {m_Tm : nat} {s0 : bool} {t0 : bool} (H0 : s0 = t0) :
-  BVal m_Tm s0 = BVal m_Tm t0.
+Lemma congr_BVal {s0 : bool} {t0 : bool} (H0 : s0 = t0) : BVal s0 = BVal t0.
 Proof.
-exact (eq_trans eq_refl (ap (fun x => BVal m_Tm x) H0)).
+exact (eq_trans eq_refl (ap (fun x => BVal x) H0)).
 Qed.
 
-Lemma congr_Bool {m_Tm : nat} : Bool m_Tm = Bool m_Tm.
+Lemma congr_Bool : Bool = Bool.
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma congr_If {m_Tm : nat} {s0 : Tm m_Tm} {s1 : Tm m_Tm} {s2 : Tm m_Tm}
-  {t0 : Tm m_Tm} {t1 : Tm m_Tm} {t2 : Tm m_Tm} (H0 : s0 = t0) (H1 : s1 = t1)
-  (H2 : s2 = t2) : If m_Tm s0 s1 s2 = If m_Tm t0 t1 t2.
+Lemma congr_If {s0 : Tm} {s1 : Tm} {s2 : Tm} {t0 : Tm} {t1 : Tm} {t2 : Tm}
+  (H0 : s0 = t0) (H1 : s1 = t1) (H2 : s2 = t2) : If s0 s1 s2 = If t0 t1 t2.
 Proof.
 exact (eq_trans
-         (eq_trans (eq_trans eq_refl (ap (fun x => If m_Tm x s1 s2) H0))
-            (ap (fun x => If m_Tm t0 x s2) H1))
-         (ap (fun x => If m_Tm t0 t1 x) H2)).
+         (eq_trans (eq_trans eq_refl (ap (fun x => If x s1 s2) H0))
+            (ap (fun x => If t0 x s2) H1)) (ap (fun x => If t0 t1 x) H2)).
 Qed.
 
-Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) :
-  fin (S m) -> fin (S n).
+Lemma upRen_Tm_Tm (xi : nat -> nat) : nat -> nat.
 Proof.
 exact (up_ren xi).
 Defined.
 
-Lemma upRen_list_Tm_Tm (p : nat) {m : nat} {n : nat} (xi : fin m -> fin n) :
-  fin (plus p m) -> fin (plus p n).
-Proof.
-exact (upRen_p p xi).
-Defined.
-
-Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
-(s : Tm m_Tm) {struct s} : Tm n_Tm :=
+Fixpoint ren_Tm (xi_Tm : nat -> nat) (s : Tm) {struct s} : Tm :=
   match s with
-  | VarTm _ s0 => VarTm n_Tm (xi_Tm s0)
-  | Abs _ s0 => Abs n_Tm (ren_Tm (upRen_Tm_Tm xi_Tm) s0)
-  | App _ s0 s1 => App n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1)
-  | Pair _ s0 s1 => Pair n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1)
-  | Proj _ s0 s1 => Proj n_Tm s0 (ren_Tm xi_Tm s1)
-  | TBind _ s0 s1 s2 =>
-      TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
-  | Univ _ s0 => Univ n_Tm s0
-  | BVal _ s0 => BVal n_Tm s0
-  | Bool _ => Bool n_Tm
-  | If _ s0 s1 s2 =>
-      If n_Tm (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) (ren_Tm xi_Tm s2)
+  | VarTm s0 => VarTm (xi_Tm s0)
+  | Abs s0 => Abs (ren_Tm (upRen_Tm_Tm xi_Tm) s0)
+  | App s0 s1 => App (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1)
+  | Pair s0 s1 => Pair (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1)
+  | Proj s0 s1 => Proj s0 (ren_Tm xi_Tm s1)
+  | TBind s0 s1 s2 =>
+      TBind s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
+  | Univ s0 => Univ s0
+  | BVal s0 => BVal s0
+  | Bool => Bool
+  | If s0 s1 s2 => If (ren_Tm xi_Tm s0) (ren_Tm xi_Tm s1) (ren_Tm xi_Tm s2)
   end.
 
-Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) :
-  fin (S m) -> Tm (S n_Tm).
+Lemma up_Tm_Tm (sigma : nat -> Tm) : nat -> Tm.
 Proof.
-exact (scons (VarTm (S n_Tm) var_zero) (funcomp (ren_Tm shift) sigma)).
+exact (scons (VarTm var_zero) (funcomp (ren_Tm shift) sigma)).
 Defined.
 
-Lemma up_list_Tm_Tm (p : nat) {m : nat} {n_Tm : nat}
-  (sigma : fin m -> Tm n_Tm) : fin (plus p m) -> Tm (plus p n_Tm).
-Proof.
-exact (scons_p p (funcomp (VarTm (plus p n_Tm)) (zero_p p))
-         (funcomp (ren_Tm (shift_p p)) sigma)).
-Defined.
-
-Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
-(s : Tm m_Tm) {struct s} : Tm n_Tm :=
+Fixpoint subst_Tm (sigma_Tm : nat -> Tm) (s : Tm) {struct s} : Tm :=
   match s with
-  | VarTm _ s0 => sigma_Tm s0
-  | Abs _ s0 => Abs n_Tm (subst_Tm (up_Tm_Tm sigma_Tm) s0)
-  | App _ s0 s1 => App n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
-  | Pair _ s0 s1 => Pair n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
-  | Proj _ s0 s1 => Proj n_Tm s0 (subst_Tm sigma_Tm s1)
-  | TBind _ s0 s1 s2 =>
-      TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
-  | Univ _ s0 => Univ n_Tm s0
-  | BVal _ s0 => BVal n_Tm s0
-  | Bool _ => Bool n_Tm
-  | If _ s0 s1 s2 =>
-      If n_Tm (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
-        (subst_Tm sigma_Tm s2)
+  | VarTm s0 => sigma_Tm s0
+  | Abs s0 => Abs (subst_Tm (up_Tm_Tm sigma_Tm) s0)
+  | App s0 s1 => App (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
+  | Pair s0 s1 => Pair (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1)
+  | Proj s0 s1 => Proj s0 (subst_Tm sigma_Tm s1)
+  | TBind s0 s1 s2 =>
+      TBind s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
+  | Univ s0 => Univ s0
+  | BVal s0 => BVal s0
+  | Bool => Bool
+  | If s0 s1 s2 =>
+      If (subst_Tm sigma_Tm s0) (subst_Tm sigma_Tm s1) (subst_Tm sigma_Tm s2)
   end.
 
-Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm)
-  (Eq : forall x, sigma x = VarTm m_Tm x) :
-  forall x, up_Tm_Tm sigma x = VarTm (S m_Tm) x.
+Lemma upId_Tm_Tm (sigma : nat -> Tm) (Eq : forall x, sigma x = VarTm x) :
+  forall x, up_Tm_Tm sigma x = VarTm x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_Tm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_Tm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma upId_list_Tm_Tm {p : nat} {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm)
-  (Eq : forall x, sigma x = VarTm m_Tm x) :
-  forall x, up_list_Tm_Tm p sigma x = VarTm (plus p m_Tm) x.
-Proof.
-exact (fun n =>
-       scons_p_eta (VarTm (plus p m_Tm))
-         (fun n => ap (ren_Tm (shift_p p)) (Eq n)) (fun n => eq_refl)).
-Qed.
-
-Fixpoint idSubst_Tm {m_Tm : nat} (sigma_Tm : fin m_Tm -> Tm m_Tm)
-(Eq_Tm : forall x, sigma_Tm x = VarTm m_Tm x) (s : Tm m_Tm) {struct s} :
+Fixpoint idSubst_Tm (sigma_Tm : nat -> Tm)
+(Eq_Tm : forall x, sigma_Tm x = VarTm x) (s : Tm) {struct s} :
 subst_Tm sigma_Tm s = s :=
   match s with
-  | VarTm _ s0 => Eq_Tm s0
-  | Abs _ s0 =>
+  | VarTm s0 => Eq_Tm s0
+  | Abs s0 =>
       congr_Abs (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (idSubst_Tm sigma_Tm Eq_Tm s0)
         (idSubst_Tm sigma_Tm Eq_Tm s1)
-  | Proj _ s0 s1 => congr_Proj (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | Proj s0 s1 => congr_Proj (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
         (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (idSubst_Tm sigma_Tm Eq_Tm s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
         (idSubst_Tm sigma_Tm Eq_Tm s2)
   end.
 
-Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n)
-  (zeta : fin m -> fin n) (Eq : forall x, xi x = zeta x) :
+Lemma upExtRen_Tm_Tm (xi : nat -> nat) (zeta : nat -> nat)
+  (Eq : forall x, xi x = zeta x) :
   forall x, upRen_Tm_Tm xi x = upRen_Tm_Tm zeta x.
 Proof.
-exact (fun n =>
-       match n with
-       | Some fin_n => ap shift (Eq fin_n)
-       | None => eq_refl
-       end).
+exact (fun n => match n with
+                | S n' => ap shift (Eq n')
+                | O => eq_refl
+                end).
 Qed.
 
-Lemma upExtRen_list_Tm_Tm {p : nat} {m : nat} {n : nat} (xi : fin m -> fin n)
-  (zeta : fin m -> fin n) (Eq : forall x, xi x = zeta x) :
-  forall x, upRen_list_Tm_Tm p xi x = upRen_list_Tm_Tm p zeta x.
-Proof.
-exact (fun n =>
-       scons_p_congr (fun n => eq_refl) (fun n => ap (shift_p p) (Eq n))).
-Qed.
-
-Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
-(zeta_Tm : fin m_Tm -> fin n_Tm) (Eq_Tm : forall x, xi_Tm x = zeta_Tm x)
-(s : Tm m_Tm) {struct s} : ren_Tm xi_Tm s = ren_Tm zeta_Tm s :=
+Fixpoint extRen_Tm (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat)
+(Eq_Tm : forall x, xi_Tm x = zeta_Tm x) (s : Tm) {struct s} :
+ren_Tm xi_Tm s = ren_Tm zeta_Tm s :=
   match s with
-  | VarTm _ s0 => ap (VarTm n_Tm) (Eq_Tm s0)
-  | Abs _ s0 =>
+  | VarTm s0 => ap (VarTm) (Eq_Tm s0)
+  | Abs s0 =>
       congr_Abs
         (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upExtRen_Tm_Tm _ _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0)
         (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0)
         (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
-  | Proj _ s0 s1 =>
-      congr_Proj (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | Proj s0 s1 => congr_Proj (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
         (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upExtRen_Tm_Tm _ _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (extRen_Tm xi_Tm zeta_Tm Eq_Tm s0)
         (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s2)
   end.
 
-Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm)
-  (tau : fin m -> Tm n_Tm) (Eq : forall x, sigma x = tau x) :
+Lemma upExt_Tm_Tm (sigma : nat -> Tm) (tau : nat -> Tm)
+  (Eq : forall x, sigma x = tau x) :
   forall x, up_Tm_Tm sigma x = up_Tm_Tm tau x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_Tm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_Tm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma upExt_list_Tm_Tm {p : nat} {m : nat} {n_Tm : nat}
-  (sigma : fin m -> Tm n_Tm) (tau : fin m -> Tm n_Tm)
-  (Eq : forall x, sigma x = tau x) :
-  forall x, up_list_Tm_Tm p sigma x = up_list_Tm_Tm p tau x.
-Proof.
-exact (fun n =>
-       scons_p_congr (fun n => eq_refl)
-         (fun n => ap (ren_Tm (shift_p p)) (Eq n))).
-Qed.
-
-Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
-(tau_Tm : fin m_Tm -> Tm n_Tm) (Eq_Tm : forall x, sigma_Tm x = tau_Tm x)
-(s : Tm m_Tm) {struct s} : subst_Tm sigma_Tm s = subst_Tm tau_Tm s :=
+Fixpoint ext_Tm (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm)
+(Eq_Tm : forall x, sigma_Tm x = tau_Tm x) (s : Tm) {struct s} :
+subst_Tm sigma_Tm s = subst_Tm tau_Tm s :=
   match s with
-  | VarTm _ s0 => Eq_Tm s0
-  | Abs _ s0 =>
+  | VarTm s0 => Eq_Tm s0
+  | Abs s0 =>
       congr_Abs
         (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm)
            s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (ext_Tm sigma_Tm tau_Tm Eq_Tm s0)
         (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (ext_Tm sigma_Tm tau_Tm Eq_Tm s0)
         (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
-  | Proj _ s0 s1 => congr_Proj (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | Proj s0 s1 => congr_Proj (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
         (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm)
            s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (ext_Tm sigma_Tm tau_Tm Eq_Tm s0)
         (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) (ext_Tm sigma_Tm tau_Tm Eq_Tm s2)
   end.
 
-Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
-  (zeta : fin l -> fin m) (rho : fin k -> fin m)
-  (Eq : forall x, funcomp zeta xi x = rho x) :
+Lemma up_ren_ren_Tm_Tm (xi : nat -> nat) (zeta : nat -> nat)
+  (rho : nat -> nat) (Eq : forall x, funcomp zeta xi x = rho x) :
   forall x, funcomp (upRen_Tm_Tm zeta) (upRen_Tm_Tm xi) x = upRen_Tm_Tm rho x.
 Proof.
 exact (up_ren_ren xi zeta rho Eq).
 Qed.
 
-Lemma up_ren_ren_list_Tm_Tm {p : nat} {k : nat} {l : nat} {m : nat}
-  (xi : fin k -> fin l) (zeta : fin l -> fin m) (rho : fin k -> fin m)
-  (Eq : forall x, funcomp zeta xi x = rho x) :
-  forall x,
-  funcomp (upRen_list_Tm_Tm p zeta) (upRen_list_Tm_Tm p xi) x =
-  upRen_list_Tm_Tm p rho x.
-Proof.
-exact (up_ren_ren_p Eq).
-Qed.
-
-Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-(xi_Tm : fin m_Tm -> fin k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm)
-(rho_Tm : fin m_Tm -> fin l_Tm)
-(Eq_Tm : forall x, funcomp zeta_Tm xi_Tm x = rho_Tm x) (s : Tm m_Tm) {struct
- s} : ren_Tm zeta_Tm (ren_Tm xi_Tm s) = ren_Tm rho_Tm s :=
+Fixpoint compRenRen_Tm (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat)
+(rho_Tm : nat -> nat) (Eq_Tm : forall x, funcomp zeta_Tm xi_Tm x = rho_Tm x)
+(s : Tm) {struct s} : ren_Tm zeta_Tm (ren_Tm xi_Tm s) = ren_Tm rho_Tm s :=
   match s with
-  | VarTm _ s0 => ap (VarTm l_Tm) (Eq_Tm s0)
-  | Abs _ s0 =>
+  | VarTm s0 => ap (VarTm) (Eq_Tm s0)
+  | Abs s0 =>
       congr_Abs
         (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0)
         (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0)
         (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
-  | Proj _ s0 s1 =>
+  | Proj s0 s1 =>
       congr_Proj (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
         (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
            (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s0)
         (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
         (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s2)
   end.
 
-Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat}
-  (xi : fin k -> fin l) (tau : fin l -> Tm m_Tm) (theta : fin k -> Tm m_Tm)
-  (Eq : forall x, funcomp tau xi x = theta x) :
+Lemma up_ren_subst_Tm_Tm (xi : nat -> nat) (tau : nat -> Tm)
+  (theta : nat -> Tm) (Eq : forall x, funcomp tau xi x = theta x) :
   forall x, funcomp (up_Tm_Tm tau) (upRen_Tm_Tm xi) x = up_Tm_Tm theta x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_Tm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_Tm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma up_ren_subst_list_Tm_Tm {p : nat} {k : nat} {l : nat} {m_Tm : nat}
-  (xi : fin k -> fin l) (tau : fin l -> Tm m_Tm) (theta : fin k -> Tm m_Tm)
-  (Eq : forall x, funcomp tau xi x = theta x) :
-  forall x,
-  funcomp (up_list_Tm_Tm p tau) (upRen_list_Tm_Tm p xi) x =
-  up_list_Tm_Tm p theta x.
-Proof.
-exact (fun n =>
-       eq_trans (scons_p_comp' _ _ _ n)
-         (scons_p_congr (fun z => scons_p_head' _ _ z)
-            (fun z =>
-             eq_trans (scons_p_tail' _ _ (xi z))
-               (ap (ren_Tm (shift_p p)) (Eq z))))).
-Qed.
-
-Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-(xi_Tm : fin m_Tm -> fin k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm)
-(theta_Tm : fin m_Tm -> Tm l_Tm)
-(Eq_Tm : forall x, funcomp tau_Tm xi_Tm x = theta_Tm x) (s : Tm m_Tm) {struct
- s} : subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm theta_Tm s :=
+Fixpoint compRenSubst_Tm (xi_Tm : nat -> nat) (tau_Tm : nat -> Tm)
+(theta_Tm : nat -> Tm)
+(Eq_Tm : forall x, funcomp tau_Tm xi_Tm x = theta_Tm x) (s : Tm) {struct s} :
+subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm theta_Tm s :=
   match s with
-  | VarTm _ s0 => Eq_Tm s0
-  | Abs _ s0 =>
+  | VarTm s0 => Eq_Tm s0
+  | Abs s0 =>
       congr_Abs
         (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0)
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0)
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
-  | Proj _ s0 s1 =>
+  | Proj s0 s1 =>
       congr_Proj (eq_refl s0)
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0)
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s0)
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s2)
   end.
 
-Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma : fin k -> Tm l_Tm) (zeta_Tm : fin l_Tm -> fin m_Tm)
-  (theta : fin k -> Tm m_Tm)
+Lemma up_subst_ren_Tm_Tm (sigma : nat -> Tm) (zeta_Tm : nat -> nat)
+  (theta : nat -> Tm)
   (Eq : forall x, funcomp (ren_Tm zeta_Tm) sigma x = theta x) :
   forall x,
   funcomp (ren_Tm (upRen_Tm_Tm zeta_Tm)) (up_Tm_Tm sigma) x =
@@ -1102,341 +840,257 @@ Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n =>
+       | S n' =>
            eq_trans
              (compRenRen_Tm shift (upRen_Tm_Tm zeta_Tm)
-                (funcomp shift zeta_Tm) (fun x => eq_refl) (sigma fin_n))
+                (funcomp shift zeta_Tm) (fun x => eq_refl) (sigma n'))
              (eq_trans
                 (eq_sym
                    (compRenRen_Tm zeta_Tm shift (funcomp shift zeta_Tm)
-                      (fun x => eq_refl) (sigma fin_n)))
-                (ap (ren_Tm shift) (Eq fin_n)))
-       | None => eq_refl
+                      (fun x => eq_refl) (sigma n')))
+                (ap (ren_Tm shift) (Eq n')))
+       | O => eq_refl
        end).
 Qed.
 
-Lemma up_subst_ren_list_Tm_Tm {p : nat} {k : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma : fin k -> Tm l_Tm) (zeta_Tm : fin l_Tm -> fin m_Tm)
-  (theta : fin k -> Tm m_Tm)
-  (Eq : forall x, funcomp (ren_Tm zeta_Tm) sigma x = theta x) :
-  forall x,
-  funcomp (ren_Tm (upRen_list_Tm_Tm p zeta_Tm)) (up_list_Tm_Tm p sigma) x =
-  up_list_Tm_Tm p theta x.
-Proof.
-exact (fun n =>
-       eq_trans (scons_p_comp' _ _ _ n)
-         (scons_p_congr
-            (fun x => ap (VarTm (plus p m_Tm)) (scons_p_head' _ _ x))
-            (fun n =>
-             eq_trans
-               (compRenRen_Tm (shift_p p) (upRen_list_Tm_Tm p zeta_Tm)
-                  (funcomp (shift_p p) zeta_Tm)
-                  (fun x => scons_p_tail' _ _ x) (sigma n))
-               (eq_trans
-                  (eq_sym
-                     (compRenRen_Tm zeta_Tm (shift_p p)
-                        (funcomp (shift_p p) zeta_Tm) (fun x => eq_refl)
-                        (sigma n))) (ap (ren_Tm (shift_p p)) (Eq n)))))).
-Qed.
-
-Fixpoint compSubstRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-(sigma_Tm : fin m_Tm -> Tm k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm)
-(theta_Tm : fin m_Tm -> Tm l_Tm)
-(Eq_Tm : forall x, funcomp (ren_Tm zeta_Tm) sigma_Tm x = theta_Tm x)
-(s : Tm m_Tm) {struct s} :
+Fixpoint compSubstRen_Tm (sigma_Tm : nat -> Tm) (zeta_Tm : nat -> nat)
+(theta_Tm : nat -> Tm)
+(Eq_Tm : forall x, funcomp (ren_Tm zeta_Tm) sigma_Tm x = theta_Tm x) 
+(s : Tm) {struct s} :
 ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
   match s with
-  | VarTm _ s0 => Eq_Tm s0
-  | Abs _ s0 =>
+  | VarTm s0 => Eq_Tm s0
+  | Abs s0 =>
       congr_Abs
         (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0)
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0)
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
-  | Proj _ s0 s1 =>
+  | Proj s0 s1 =>
       congr_Proj (eq_refl s0)
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0)
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
         (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s0)
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
         (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s2)
   end.
 
-Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma : fin k -> Tm l_Tm) (tau_Tm : fin l_Tm -> Tm m_Tm)
-  (theta : fin k -> Tm m_Tm)
+Lemma up_subst_subst_Tm_Tm (sigma : nat -> Tm) (tau_Tm : nat -> Tm)
+  (theta : nat -> Tm)
   (Eq : forall x, funcomp (subst_Tm tau_Tm) sigma x = theta x) :
   forall x,
   funcomp (subst_Tm (up_Tm_Tm tau_Tm)) (up_Tm_Tm sigma) x = up_Tm_Tm theta x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n =>
+       | S n' =>
            eq_trans
              (compRenSubst_Tm shift (up_Tm_Tm tau_Tm)
                 (funcomp (up_Tm_Tm tau_Tm) shift) (fun x => eq_refl)
-                (sigma fin_n))
+                (sigma n'))
              (eq_trans
                 (eq_sym
                    (compSubstRen_Tm tau_Tm shift
                       (funcomp (ren_Tm shift) tau_Tm) (fun x => eq_refl)
-                      (sigma fin_n))) (ap (ren_Tm shift) (Eq fin_n)))
-       | None => eq_refl
+                      (sigma n'))) (ap (ren_Tm shift) (Eq n')))
+       | O => eq_refl
        end).
 Qed.
 
-Lemma up_subst_subst_list_Tm_Tm {p : nat} {k : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma : fin k -> Tm l_Tm) (tau_Tm : fin l_Tm -> Tm m_Tm)
-  (theta : fin k -> Tm m_Tm)
-  (Eq : forall x, funcomp (subst_Tm tau_Tm) sigma x = theta x) :
-  forall x,
-  funcomp (subst_Tm (up_list_Tm_Tm p tau_Tm)) (up_list_Tm_Tm p sigma) x =
-  up_list_Tm_Tm p theta x.
-Proof.
-exact (fun n =>
-       eq_trans
-         (scons_p_comp' (funcomp (VarTm (plus p l_Tm)) (zero_p p)) _ _ n)
-         (scons_p_congr
-            (fun x => scons_p_head' _ (fun z => ren_Tm (shift_p p) _) x)
-            (fun n =>
-             eq_trans
-               (compRenSubst_Tm (shift_p p) (up_list_Tm_Tm p tau_Tm)
-                  (funcomp (up_list_Tm_Tm p tau_Tm) (shift_p p))
-                  (fun x => eq_refl) (sigma n))
-               (eq_trans
-                  (eq_sym
-                     (compSubstRen_Tm tau_Tm (shift_p p) _
-                        (fun x => eq_sym (scons_p_tail' _ _ x)) (sigma n)))
-                  (ap (ren_Tm (shift_p p)) (Eq n)))))).
-Qed.
-
-Fixpoint compSubstSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-(sigma_Tm : fin m_Tm -> Tm k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm)
-(theta_Tm : fin m_Tm -> Tm l_Tm)
+Fixpoint compSubstSubst_Tm (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm)
+(theta_Tm : nat -> Tm)
 (Eq_Tm : forall x, funcomp (subst_Tm tau_Tm) sigma_Tm x = theta_Tm x)
-(s : Tm m_Tm) {struct s} :
+(s : Tm) {struct s} :
 subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
   match s with
-  | VarTm _ s0 => Eq_Tm s0
-  | Abs _ s0 =>
+  | VarTm s0 => Eq_Tm s0
+  | Abs s0 =>
       congr_Abs
         (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0)
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0)
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
-  | Proj _ s0 s1 =>
+  | Proj s0 s1 =>
       congr_Proj (eq_refl s0)
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0)
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm)
            (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s0)
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
         (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s2)
   end.
 
-Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (xi_Tm : fin m_Tm -> fin k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm)
-  (s : Tm m_Tm) :
+Lemma renRen_Tm (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat) (s : Tm) :
   ren_Tm zeta_Tm (ren_Tm xi_Tm s) = ren_Tm (funcomp zeta_Tm xi_Tm) s.
 Proof.
 exact (compRenRen_Tm xi_Tm zeta_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma renRen'_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (xi_Tm : fin m_Tm -> fin k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) :
+Lemma renRen'_Tm_pointwise (xi_Tm : nat -> nat) (zeta_Tm : nat -> nat) :
   pointwise_relation _ eq (funcomp (ren_Tm zeta_Tm) (ren_Tm xi_Tm))
     (ren_Tm (funcomp zeta_Tm xi_Tm)).
 Proof.
 exact (fun s => compRenRen_Tm xi_Tm zeta_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma renSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (xi_Tm : fin m_Tm -> fin k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) (s : Tm m_Tm)
-  : subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm (funcomp tau_Tm xi_Tm) s.
+Lemma renSubst_Tm (xi_Tm : nat -> nat) (tau_Tm : nat -> Tm) (s : Tm) :
+  subst_Tm tau_Tm (ren_Tm xi_Tm s) = subst_Tm (funcomp tau_Tm xi_Tm) s.
 Proof.
 exact (compRenSubst_Tm xi_Tm tau_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma renSubst_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (xi_Tm : fin m_Tm -> fin k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) :
+Lemma renSubst_Tm_pointwise (xi_Tm : nat -> nat) (tau_Tm : nat -> Tm) :
   pointwise_relation _ eq (funcomp (subst_Tm tau_Tm) (ren_Tm xi_Tm))
     (subst_Tm (funcomp tau_Tm xi_Tm)).
 Proof.
 exact (fun s => compRenSubst_Tm xi_Tm tau_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma_Tm : fin m_Tm -> Tm k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm)
-  (s : Tm m_Tm) :
+Lemma substRen_Tm (sigma_Tm : nat -> Tm) (zeta_Tm : nat -> nat) (s : Tm) :
   ren_Tm zeta_Tm (subst_Tm sigma_Tm s) =
   subst_Tm (funcomp (ren_Tm zeta_Tm) sigma_Tm) s.
 Proof.
 exact (compSubstRen_Tm sigma_Tm zeta_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substRen_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma_Tm : fin m_Tm -> Tm k_Tm) (zeta_Tm : fin k_Tm -> fin l_Tm) :
+Lemma substRen_Tm_pointwise (sigma_Tm : nat -> Tm) (zeta_Tm : nat -> nat) :
   pointwise_relation _ eq (funcomp (ren_Tm zeta_Tm) (subst_Tm sigma_Tm))
     (subst_Tm (funcomp (ren_Tm zeta_Tm) sigma_Tm)).
 Proof.
 exact (fun s => compSubstRen_Tm sigma_Tm zeta_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma_Tm : fin m_Tm -> Tm k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm)
-  (s : Tm m_Tm) :
+Lemma substSubst_Tm (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm) (s : Tm) :
   subst_Tm tau_Tm (subst_Tm sigma_Tm s) =
   subst_Tm (funcomp (subst_Tm tau_Tm) sigma_Tm) s.
 Proof.
 exact (compSubstSubst_Tm sigma_Tm tau_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma substSubst_Tm_pointwise {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
-  (sigma_Tm : fin m_Tm -> Tm k_Tm) (tau_Tm : fin k_Tm -> Tm l_Tm) :
+Lemma substSubst_Tm_pointwise (sigma_Tm : nat -> Tm) (tau_Tm : nat -> Tm) :
   pointwise_relation _ eq (funcomp (subst_Tm tau_Tm) (subst_Tm sigma_Tm))
     (subst_Tm (funcomp (subst_Tm tau_Tm) sigma_Tm)).
 Proof.
 exact (fun s => compSubstSubst_Tm sigma_Tm tau_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma rinstInst_up_Tm_Tm {m : nat} {n_Tm : nat} (xi : fin m -> fin n_Tm)
-  (sigma : fin m -> Tm n_Tm)
-  (Eq : forall x, funcomp (VarTm n_Tm) xi x = sigma x) :
-  forall x, funcomp (VarTm (S n_Tm)) (upRen_Tm_Tm xi) x = up_Tm_Tm sigma x.
+Lemma rinstInst_up_Tm_Tm (xi : nat -> nat) (sigma : nat -> Tm)
+  (Eq : forall x, funcomp (VarTm) xi x = sigma x) :
+  forall x, funcomp (VarTm) (upRen_Tm_Tm xi) x = up_Tm_Tm sigma x.
 Proof.
 exact (fun n =>
        match n with
-       | Some fin_n => ap (ren_Tm shift) (Eq fin_n)
-       | None => eq_refl
+       | S n' => ap (ren_Tm shift) (Eq n')
+       | O => eq_refl
        end).
 Qed.
 
-Lemma rinstInst_up_list_Tm_Tm {p : nat} {m : nat} {n_Tm : nat}
-  (xi : fin m -> fin n_Tm) (sigma : fin m -> Tm n_Tm)
-  (Eq : forall x, funcomp (VarTm n_Tm) xi x = sigma x) :
-  forall x,
-  funcomp (VarTm (plus p n_Tm)) (upRen_list_Tm_Tm p xi) x =
-  up_list_Tm_Tm p sigma x.
-Proof.
-exact (fun n =>
-       eq_trans (scons_p_comp' _ _ (VarTm (plus p n_Tm)) n)
-         (scons_p_congr (fun z => eq_refl)
-            (fun n => ap (ren_Tm (shift_p p)) (Eq n)))).
-Qed.
-
-Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
-(xi_Tm : fin m_Tm -> fin n_Tm) (sigma_Tm : fin m_Tm -> Tm n_Tm)
-(Eq_Tm : forall x, funcomp (VarTm n_Tm) xi_Tm x = sigma_Tm x) (s : Tm m_Tm)
-{struct s} : ren_Tm xi_Tm s = subst_Tm sigma_Tm s :=
+Fixpoint rinst_inst_Tm (xi_Tm : nat -> nat) (sigma_Tm : nat -> Tm)
+(Eq_Tm : forall x, funcomp (VarTm) xi_Tm x = sigma_Tm x) (s : Tm) {struct s}
+   : ren_Tm xi_Tm s = subst_Tm sigma_Tm s :=
   match s with
-  | VarTm _ s0 => Eq_Tm s0
-  | Abs _ s0 =>
+  | VarTm s0 => Eq_Tm s0
+  | Abs s0 =>
       congr_Abs
         (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm)
            (rinstInst_up_Tm_Tm _ _ Eq_Tm) s0)
-  | App _ s0 s1 =>
+  | App s0 s1 =>
       congr_App (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0)
         (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
-  | Pair _ s0 s1 =>
+  | Pair s0 s1 =>
       congr_Pair (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0)
         (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
-  | Proj _ s0 s1 =>
+  | Proj s0 s1 =>
       congr_Proj (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
-  | TBind _ s0 s1 s2 =>
+  | TBind s0 s1 s2 =>
       congr_TBind (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
         (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm)
            (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2)
-  | Univ _ s0 => congr_Univ (eq_refl s0)
-  | BVal _ s0 => congr_BVal (eq_refl s0)
-  | Bool _ => congr_Bool
-  | If _ s0 s1 s2 =>
+  | Univ s0 => congr_Univ (eq_refl s0)
+  | BVal s0 => congr_BVal (eq_refl s0)
+  | Bool => congr_Bool
+  | If s0 s1 s2 =>
       congr_If (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s0)
         (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
         (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s2)
   end.
 
-Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
-  (s : Tm m_Tm) : ren_Tm xi_Tm s = subst_Tm (funcomp (VarTm n_Tm) xi_Tm) s.
+Lemma rinstInst'_Tm (xi_Tm : nat -> nat) (s : Tm) :
+  ren_Tm xi_Tm s = subst_Tm (funcomp (VarTm) xi_Tm) s.
 Proof.
 exact (rinst_inst_Tm xi_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma rinstInst'_Tm_pointwise {m_Tm : nat} {n_Tm : nat}
-  (xi_Tm : fin m_Tm -> fin n_Tm) :
-  pointwise_relation _ eq (ren_Tm xi_Tm)
-    (subst_Tm (funcomp (VarTm n_Tm) xi_Tm)).
+Lemma rinstInst'_Tm_pointwise (xi_Tm : nat -> nat) :
+  pointwise_relation _ eq (ren_Tm xi_Tm) (subst_Tm (funcomp (VarTm) xi_Tm)).
 Proof.
 exact (fun s => rinst_inst_Tm xi_Tm _ (fun n => eq_refl) s).
 Qed.
 
-Lemma instId'_Tm {m_Tm : nat} (s : Tm m_Tm) : subst_Tm (VarTm m_Tm) s = s.
+Lemma instId'_Tm (s : Tm) : subst_Tm (VarTm) s = s.
 Proof.
-exact (idSubst_Tm (VarTm m_Tm) (fun n => eq_refl) s).
+exact (idSubst_Tm (VarTm) (fun n => eq_refl) s).
 Qed.
 
-Lemma instId'_Tm_pointwise {m_Tm : nat} :
-  pointwise_relation _ eq (subst_Tm (VarTm m_Tm)) id.
+Lemma instId'_Tm_pointwise : pointwise_relation _ eq (subst_Tm (VarTm)) id.
 Proof.
-exact (fun s => idSubst_Tm (VarTm m_Tm) (fun n => eq_refl) s).
+exact (fun s => idSubst_Tm (VarTm) (fun n => eq_refl) s).
 Qed.
 
-Lemma rinstId'_Tm {m_Tm : nat} (s : Tm m_Tm) : ren_Tm id s = s.
+Lemma rinstId'_Tm (s : Tm) : ren_Tm id s = s.
 Proof.
 exact (eq_ind_r (fun t => t = s) (instId'_Tm s) (rinstInst'_Tm id s)).
 Qed.
 
-Lemma rinstId'_Tm_pointwise {m_Tm : nat} :
-  pointwise_relation _ eq (@ren_Tm m_Tm m_Tm id) id.
+Lemma rinstId'_Tm_pointwise : pointwise_relation _ eq (@ren_Tm id) id.
 Proof.
 exact (fun s => eq_ind_r (fun t => t = s) (instId'_Tm s) (rinstInst'_Tm id s)).
 Qed.
 
-Lemma varL'_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
-  (x : fin m_Tm) : subst_Tm sigma_Tm (VarTm m_Tm x) = sigma_Tm x.
+Lemma varL'_Tm (sigma_Tm : nat -> Tm) (x : nat) :
+  subst_Tm sigma_Tm (VarTm x) = sigma_Tm x.
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma varL'_Tm_pointwise {m_Tm : nat} {n_Tm : nat}
-  (sigma_Tm : fin m_Tm -> Tm n_Tm) :
-  pointwise_relation _ eq (funcomp (subst_Tm sigma_Tm) (VarTm m_Tm)) sigma_Tm.
+Lemma varL'_Tm_pointwise (sigma_Tm : nat -> Tm) :
+  pointwise_relation _ eq (funcomp (subst_Tm sigma_Tm) (VarTm)) sigma_Tm.
 Proof.
 exact (fun x => eq_refl).
 Qed.
 
-Lemma varLRen'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
-  (x : fin m_Tm) : ren_Tm xi_Tm (VarTm m_Tm x) = VarTm n_Tm (xi_Tm x).
+Lemma varLRen'_Tm (xi_Tm : nat -> nat) (x : nat) :
+  ren_Tm xi_Tm (VarTm x) = VarTm (xi_Tm x).
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma varLRen'_Tm_pointwise {m_Tm : nat} {n_Tm : nat}
-  (xi_Tm : fin m_Tm -> fin n_Tm) :
-  pointwise_relation _ eq (funcomp (ren_Tm xi_Tm) (VarTm m_Tm))
-    (funcomp (VarTm n_Tm) xi_Tm).
+Lemma varLRen'_Tm_pointwise (xi_Tm : nat -> nat) :
+  pointwise_relation _ eq (funcomp (ren_Tm xi_Tm) (VarTm))
+    (funcomp (VarTm) xi_Tm).
 Proof.
 exact (fun x => eq_refl).
 Qed.
@@ -1447,30 +1101,22 @@ Class Up_Tm X Y :=
 Class Up_PTm X Y :=
     up_PTm : X -> Y.
 
-#[global]
-Instance Subst_Tm  {m_Tm n_Tm : nat}: (Subst1 _ _ _) := (@subst_Tm m_Tm n_Tm).
+#[global] Instance Subst_Tm : (Subst1 _ _ _) := @subst_Tm.
+
+#[global] Instance Up_Tm_Tm : (Up_Tm _ _) := @up_Tm_Tm.
+
+#[global] Instance Ren_Tm : (Ren1 _ _ _) := @ren_Tm.
+
+#[global] Instance VarInstance_Tm : (Var _ _) := @VarTm.
+
+#[global] Instance Subst_PTm : (Subst1 _ _ _) := @subst_PTm.
+
+#[global] Instance Up_PTm_PTm : (Up_PTm _ _) := @up_PTm_PTm.
+
+#[global] Instance Ren_PTm : (Ren1 _ _ _) := @ren_PTm.
 
 #[global]
-Instance Up_Tm_Tm  {m n_Tm : nat}: (Up_Tm _ _) := (@up_Tm_Tm m n_Tm).
-
-#[global]
-Instance Ren_Tm  {m_Tm n_Tm : nat}: (Ren1 _ _ _) := (@ren_Tm m_Tm n_Tm).
-
-#[global]
-Instance VarInstance_Tm  {n_Tm : nat}: (Var _ _) := (@VarTm n_Tm).
-
-#[global]
-Instance Subst_PTm  {m_PTm n_PTm : nat}: (Subst1 _ _ _) :=
- (@subst_PTm m_PTm n_PTm).
-
-#[global]
-Instance Up_PTm_PTm  {m n_PTm : nat}: (Up_PTm _ _) := (@up_PTm_PTm m n_PTm).
-
-#[global]
-Instance Ren_PTm  {m_PTm n_PTm : nat}: (Ren1 _ _ _) := (@ren_PTm m_PTm n_PTm).
-
-#[global]
-Instance VarInstance_PTm  {n_PTm : nat}: (Var _ _) := (@VarPTm n_PTm).
+Instance VarInstance_PTm : (Var _ _) := @VarPTm.
 
 Notation "[ sigma_Tm ]" := (subst_Tm sigma_Tm)
 ( at level 1, left associativity, only printing)  : fscope.
@@ -1521,9 +1167,9 @@ Notation "x '__PTm'" := (VarPTm x) ( at level 5, format "x __PTm")  :
 subst_scope.
 
 #[global]
-Instance subst_Tm_morphism  {m_Tm : nat} {n_Tm : nat}:
+Instance subst_Tm_morphism :
  (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
-    (@subst_Tm m_Tm n_Tm)).
+    (@subst_Tm)).
 Proof.
 exact (fun f_Tm g_Tm Eq_Tm s t Eq_st =>
        eq_ind s (fun t' => subst_Tm f_Tm s = subst_Tm g_Tm t')
@@ -1531,17 +1177,16 @@ exact (fun f_Tm g_Tm Eq_Tm s t Eq_st =>
 Qed.
 
 #[global]
-Instance subst_Tm_morphism2  {m_Tm : nat} {n_Tm : nat}:
+Instance subst_Tm_morphism2 :
  (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq))
-    (@subst_Tm m_Tm n_Tm)).
+    (@subst_Tm)).
 Proof.
 exact (fun f_Tm g_Tm Eq_Tm s => ext_Tm f_Tm g_Tm Eq_Tm s).
 Qed.
 
 #[global]
-Instance ren_Tm_morphism  {m_Tm : nat} {n_Tm : nat}:
- (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
-    (@ren_Tm m_Tm n_Tm)).
+Instance ren_Tm_morphism :
+ (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) (@ren_Tm)).
 Proof.
 exact (fun f_Tm g_Tm Eq_Tm s t Eq_st =>
        eq_ind s (fun t' => ren_Tm f_Tm s = ren_Tm g_Tm t')
@@ -1549,17 +1194,17 @@ exact (fun f_Tm g_Tm Eq_Tm s t Eq_st =>
 Qed.
 
 #[global]
-Instance ren_Tm_morphism2  {m_Tm : nat} {n_Tm : nat}:
+Instance ren_Tm_morphism2 :
  (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq))
-    (@ren_Tm m_Tm n_Tm)).
+    (@ren_Tm)).
 Proof.
 exact (fun f_Tm g_Tm Eq_Tm s => extRen_Tm f_Tm g_Tm Eq_Tm s).
 Qed.
 
 #[global]
-Instance subst_PTm_morphism  {m_PTm : nat} {n_PTm : nat}:
+Instance subst_PTm_morphism :
  (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
-    (@subst_PTm m_PTm n_PTm)).
+    (@subst_PTm)).
 Proof.
 exact (fun f_PTm g_PTm Eq_PTm s t Eq_st =>
        eq_ind s (fun t' => subst_PTm f_PTm s = subst_PTm g_PTm t')
@@ -1567,17 +1212,16 @@ exact (fun f_PTm g_PTm Eq_PTm s t Eq_st =>
 Qed.
 
 #[global]
-Instance subst_PTm_morphism2  {m_PTm : nat} {n_PTm : nat}:
+Instance subst_PTm_morphism2 :
  (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq))
-    (@subst_PTm m_PTm n_PTm)).
+    (@subst_PTm)).
 Proof.
 exact (fun f_PTm g_PTm Eq_PTm s => ext_PTm f_PTm g_PTm Eq_PTm s).
 Qed.
 
 #[global]
-Instance ren_PTm_morphism  {m_PTm : nat} {n_PTm : nat}:
- (Proper (respectful (pointwise_relation _ eq) (respectful eq eq))
-    (@ren_PTm m_PTm n_PTm)).
+Instance ren_PTm_morphism :
+ (Proper (respectful (pointwise_relation _ eq) (respectful eq eq)) (@ren_PTm)).
 Proof.
 exact (fun f_PTm g_PTm Eq_PTm s t Eq_st =>
        eq_ind s (fun t' => ren_PTm f_PTm s = ren_PTm g_PTm t')
@@ -1585,9 +1229,9 @@ exact (fun f_PTm g_PTm Eq_PTm s t Eq_st =>
 Qed.
 
 #[global]
-Instance ren_PTm_morphism2  {m_PTm : nat} {n_PTm : nat}:
+Instance ren_PTm_morphism2 :
  (Proper (respectful (pointwise_relation _ eq) (pointwise_relation _ eq))
-    (@ren_PTm m_PTm n_PTm)).
+    (@ren_PTm)).
 Proof.
 exact (fun f_PTm g_PTm Eq_PTm s => extRen_PTm f_PTm g_PTm Eq_PTm s).
 Qed.
@@ -1641,9 +1285,8 @@ Ltac asimpl' := repeat (first
                  | progress setoid_rewrite instId'_PTm_pointwise
                  | progress setoid_rewrite instId'_PTm
                  | progress
-                    unfold up_list_Tm_Tm, up_Tm_Tm, upRen_list_Tm_Tm,
-                     upRen_Tm_Tm, up_list_PTm_PTm, up_PTm_PTm,
-                     upRen_list_PTm_PTm, upRen_PTm_PTm, up_ren
+                    unfold up_Tm_Tm, upRen_Tm_Tm, up_PTm_PTm, upRen_PTm_PTm,
+                     up_ren
                  | progress cbn[subst_Tm ren_Tm subst_PTm ren_PTm]
                  | progress fsimpl ]).
 
@@ -1673,52 +1316,15 @@ End Core.
 
 Module Extra.
 
-Import
-Core.
+Import Core.
 
-Arguments VarTm {n_Tm}.
+#[global] Hint Opaque subst_Tm: rewrite.
 
-Arguments If {n_Tm}.
+#[global] Hint Opaque ren_Tm: rewrite.
 
-Arguments Bool {n_Tm}.
+#[global] Hint Opaque subst_PTm: rewrite.
 
-Arguments BVal {n_Tm}.
-
-Arguments Univ {n_Tm}.
-
-Arguments TBind {n_Tm}.
-
-Arguments Proj {n_Tm}.
-
-Arguments Pair {n_Tm}.
-
-Arguments App {n_Tm}.
-
-Arguments Abs {n_Tm}.
-
-Arguments VarPTm {n_PTm}.
-
-Arguments PBot {n_PTm}.
-
-Arguments PUniv {n_PTm}.
-
-Arguments PConst {n_PTm}.
-
-Arguments PProj {n_PTm}.
-
-Arguments PPair {n_PTm}.
-
-Arguments PApp {n_PTm}.
-
-Arguments PAbs {n_PTm}.
-
-#[global]Hint Opaque subst_Tm: rewrite.
-
-#[global]Hint Opaque ren_Tm: rewrite.
-
-#[global]Hint Opaque subst_PTm: rewrite.
-
-#[global]Hint Opaque ren_PTm: rewrite.
+#[global] Hint Opaque ren_PTm: rewrite.
 
 End Extra.
 
diff --git a/theories/Autosubst2/unscoped.v b/theories/Autosubst2/unscoped.v
new file mode 100644
index 0000000..55f8172
--- /dev/null
+++ b/theories/Autosubst2/unscoped.v
@@ -0,0 +1,213 @@
+(** * Autosubst Header for Unnamed Syntax
+
+Version: December 11, 2019.
+ *)
+
+(* Adrian:
+ I changed this library a bit to work better with my generated code.
+ 1. I use nat directly instead of defining fin to be nat and using Some/None as S/O
+ 2. I removed the "s, sigma" notation for scons because it interacts with dependent function types "forall x, A"*)
+Require Import core.
+Require Import Setoid Morphisms Relation_Definitions.
+
+Definition ap {X Y} (f : X -> Y) {x y : X} (p : x = y) : f x = f y :=
+  match p with eq_refl => eq_refl end.
+
+Definition apc {X Y} {f g : X -> Y} {x y : X} (p : f = g) (q : x = y) : f x = g y :=
+  match q with eq_refl => match p with eq_refl => eq_refl end end.
+
+(** ** Primitives of the Sigma Calculus. *)
+
+Definition shift  := S.
+
+Definition var_zero := 0.
+
+Definition id {X} := @Datatypes.id X.
+
+Definition scons {X: Type} (x : X) (xi : nat -> X) :=
+  fun n => match n with
+        | 0 => x
+        | S n => xi n
+        end.
+
+#[ export ]
+Hint Opaque scons : rewrite.
+
+(** ** Type Class Instances for Notation
+Required to make notation work. *)
+
+(** *** Type classes for renamings. *)
+
+Class Ren1 (X1  : Type) (Y Z : Type) :=
+  ren1 : X1 -> Y -> Z.
+
+Class Ren2 (X1 X2 : Type) (Y Z : Type) :=
+  ren2 : X1 -> X2 -> Y -> Z.
+
+Class Ren3 (X1 X2 X3 : Type) (Y Z : Type) :=
+  ren3 : X1 -> X2 -> X3 -> Y -> Z.
+
+Class Ren4 (X1 X2 X3 X4 : Type) (Y Z : Type) :=
+  ren4 : X1 -> X2 -> X3 -> X4 -> Y -> Z.
+
+Class Ren5 (X1 X2 X3 X4 X5 : Type) (Y Z : Type) :=
+  ren5 : X1 -> X2 -> X3 -> X4 -> X5 -> Y -> Z.
+
+Module RenNotations.
+  Notation "s ⟨ xi1 ⟩" := (ren1 xi1 s) (at level 7, left associativity, format "s  ⟨ xi1 ⟩") : subst_scope.
+
+  Notation "s ⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ⟩") : subst_scope.
+
+  Notation "s ⟨ xi1 ; xi2 ; xi3 ⟩" := (ren3 xi1 xi2 xi3 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ; xi3 ⟩") : subst_scope.
+
+  Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩" := (ren4  xi1 xi2 xi3 xi4 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ; xi3 ; xi4 ⟩") : subst_scope.
+
+  Notation "s ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩" := (ren5  xi1 xi2 xi3 xi4 xi5 s) (at level 7, left associativity, format "s  ⟨ xi1 ; xi2 ; xi3 ; xi4 ; xi5 ⟩") : subst_scope.
+
+  Notation "⟨ xi ⟩" := (ren1 xi) (at level 1, left associativity, format "⟨ xi ⟩") : fscope.
+
+  Notation "⟨ xi1 ; xi2 ⟩" := (ren2 xi1 xi2) (at level 1, left associativity, format "⟨ xi1 ; xi2 ⟩") : fscope.
+End RenNotations.
+
+(** *** Type Classes for Substiution *)
+
+Class Subst1 (X1 : Type) (Y Z: Type) :=
+  subst1 : X1 -> Y -> Z.
+
+Class Subst2 (X1 X2 : Type) (Y Z: Type) :=
+  subst2 : X1 -> X2 -> Y  -> Z.
+
+Class Subst3 (X1 X2 X3 : Type) (Y Z: Type) :=
+  subst3 : X1 -> X2 -> X3 ->  Y  -> Z.
+
+Class Subst4 (X1 X2 X3 X4: Type) (Y Z: Type) :=
+  subst4 : X1 -> X2 -> X3 -> X4 -> Y  -> Z.
+
+Class Subst5 (X1 X2 X3 X4 X5 : Type) (Y Z: Type) :=
+  subst5 : X1 -> X2 -> X3 -> X4 -> X5  -> Y  -> Z.
+
+Module SubstNotations.
+  Notation "s [ sigma ]" := (subst1 sigma s) (at level 7, left associativity, format "s '/' [ sigma ]") : subst_scope.
+
+  Notation "s [ sigma ; tau ]" := (subst2 sigma tau s) (at level 7, left associativity, format "s '/' [ sigma ; '/'  tau ]") : subst_scope.
+End SubstNotations.
+
+(** *** Type Class for Variables *)
+
+Class Var X Y :=
+  ids : X -> Y.
+
+#[export] Instance idsRen : Var nat nat := id.
+
+(** ** Proofs for the substitution primitives. *)
+
+Arguments funcomp {X Y Z} (g)%fscope (f)%fscope.
+
+Module CombineNotations.
+  Notation "f >> g" := (funcomp g f) (at level 50) : fscope.
+
+  Notation "s .: sigma" := (scons s sigma) (at level 55, sigma at next level, right associativity) : subst_scope.
+
+  #[ global ]
+  Open Scope fscope.
+  #[ global ]
+  Open Scope subst_scope.
+End CombineNotations.
+
+Import CombineNotations.
+
+
+(** A generic lifting of a renaming. *)
+Definition up_ren (xi : nat -> nat) :=
+  0 .: (xi >> S).
+
+(** A generic proof that lifting of renamings composes. *)
+Lemma up_ren_ren (xi: nat -> nat) (zeta : nat -> nat) (rho: nat -> nat) (E: forall x, (xi >> zeta) x = rho x) :
+  forall x, (up_ren xi >> up_ren zeta) x = up_ren rho x.
+Proof.
+  intros [|x].
+  - reflexivity.
+  - unfold up_ren. cbn. unfold funcomp. f_equal. apply E.
+Qed.
+
+(** Eta laws. *)
+Lemma scons_eta' {T} (f : nat -> T) :
+  pointwise_relation _ eq (f var_zero .: (funcomp f shift)) f.
+Proof. intros x. destruct x; reflexivity. Qed.
+
+Lemma scons_eta_id' :
+  pointwise_relation _ eq (var_zero .: shift) id.
+Proof. intros x. destruct x; reflexivity. Qed.
+
+Lemma scons_comp' (T: Type) {U} (s: T) (sigma: nat -> T) (tau: T -> U) :
+  pointwise_relation _ eq (funcomp tau (s .: sigma)) ((tau s) .: (funcomp tau sigma)).
+Proof. intros x. destruct x; reflexivity. Qed.
+
+(* Morphism for Setoid Rewriting. The only morphism that can be defined statically. *)
+#[export] Instance scons_morphism {X: Type} :
+  Proper (eq ==> pointwise_relation _ eq ==> pointwise_relation _ eq) (@scons X).
+Proof.
+  intros ? t -> sigma tau H.
+  intros [|x].
+  cbn. reflexivity.
+  apply H.
+Qed.
+
+#[export] Instance scons_morphism2 {X: Type} :
+  Proper (eq ==> pointwise_relation _ eq ==> eq ==> eq) (@scons X).
+Proof.
+  intros ? t -> sigma tau H ? x ->.
+  destruct x as [|x].
+  cbn. reflexivity.
+  apply H.
+Qed.
+
+(** ** Generic lifting of an allfv predicate *)
+Definition up_allfv (p: nat -> Prop) : nat -> Prop := scons True p.
+
+(** ** Notations for unscoped syntax *)
+Module UnscopedNotations.
+  Include RenNotations.
+  Include SubstNotations.
+  Include CombineNotations.
+  
+  (* Notation "s , sigma" := (scons s sigma) (at level 60, format "s ,  sigma", right associativity) : subst_scope. *)
+
+  Notation "s '..'" := (scons s ids) (at level 1, format "s ..") : subst_scope.
+
+  Notation "↑" := (shift) : subst_scope.
+
+  #[global]
+  Open Scope fscope.
+  #[global]
+  Open Scope subst_scope.
+End UnscopedNotations.
+
+(** ** Tactics for unscoped syntax *)
+
+(** Automatically does a case analysis on a natural number, useful for proofs with context renamings/context morphisms. *)
+Tactic Notation "auto_case" tactic(t) :=  (match goal with
+                                           | [|- forall (i : nat), _] => intros []; t
+                                           end).
+
+
+(** Generic fsimpl tactic: simplifies the above primitives in a goal. *)
+Ltac fsimpl :=
+  repeat match goal with
+         | [|- context[id >> ?f]] => change (id >> f) with f (* AsimplCompIdL *)
+         | [|- context[?f >> id]] => change (f >> id) with f (* AsimplCompIdR *)
+         | [|- context [id ?s]] => change (id s) with s
+         | [|- context[(?f >> ?g) >> ?h]] => change ((f >> g) >> h) with (f >> (g >> h))
+         | [|- context[(?v .: ?g) var_zero]] => change ((v .: g) var_zero) with v
+         | [|- context[(?v .: ?g) 0]] => change ((v .: g) 0) with v
+         | [|- context[(?v .: ?g) (S ?n)]] => change ((v .: g) (S n)) with (g n)
+         | [|- context[?f >> (?x .: ?g)]] => change (f >> (x .: g)) with g (* f should evaluate to shift *)
+         | [|- context[var_zero]] =>  change var_zero with 0
+         | [|- context[?x2 .: (funcomp ?f shift)]] => change (scons x2 (funcomp f shift)) with (scons (f var_zero) (funcomp f shift)); setoid_rewrite (@scons_eta' _ _ f)
+         | [|- context[?f var_zero .: ?g]] => change (scons (f var_zero) g) with (scons (f var_zero) (funcomp f shift)); rewrite scons_eta'
+         | [|- _ =  ?h (?f ?s)] => change (h (f s)) with ((f >> h) s)
+         | [|-  ?h (?f ?s) = _] => change (h (f s)) with ((f >> h) s)
+         (* DONE had to put an underscore as the last argument to scons. This might be an argument against unfolding funcomp *)
+         | [|- context[funcomp _ (scons _ _)]] => setoid_rewrite scons_comp'; eta_reduce
+         | [|- context[scons var_zero shift]] => setoid_rewrite scons_eta_id'; eta_reduce
+                        end.
\ No newline at end of file
diff --git a/theories/fp_red.v b/theories/fp_red.v
index bbe58d9..f9abc08 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -7,7 +7,7 @@ Require Import Arith.Wf_nat.
 Require Import Psatz.
 From stdpp Require Import relations (rtc (..), rtc_once, rtc_r).
 From Hammer Require Import Tactics.
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax.
 
 Ltac2 spec_refl () :=
   List.iter
@@ -22,7 +22,7 @@ Ltac spec_refl := ltac2:(spec_refl ()).
 
 (* Trying my best to not write C style module_funcname *)
 Module Par.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (***************** Beta ***********************)
   | AppAbs a0 a1 b0 b1 :
     R a0 a1 ->
@@ -72,35 +72,35 @@ Module Par.
   | Bot :
     R PBot PBot.
 
-  Lemma refl n (a : PTm n) : R a a.
-    elim : n /a; hauto ctrs:R.
+  Lemma refl (a : PTm) : R a a.
+    elim : a; hauto ctrs:R.
   Qed.
 
-  Lemma AppAbs' n a0 a1 (b0 b1 t : PTm n) :
+  Lemma AppAbs' a0 a1 (b0 b1 t : PTm) :
     t = subst_PTm (scons b1 VarPTm) a1 ->
     R a0 a1 ->
     R b0 b1 ->
     R (PApp (PAbs a0) b0) t.
   Proof. move => ->. apply AppAbs. Qed.
 
-  Lemma ProjPair' n p (a0 a1 b0 b1 : PTm n) t :
+  Lemma ProjPair' p (a0 a1 b0 b1 : PTm) t :
     t = (if p is PL then a1 else b1) ->
     R a0 a1 ->
     R b0 b1 ->
     R (PProj p (PPair a0 b0)) t.
   Proof.  move => > ->. apply ProjPair. Qed.
 
-  Lemma AppEta' n (a0 a1 b : PTm n) :
+  Lemma AppEta' (a0 a1 b : PTm) :
     b = (PAbs (PApp (ren_PTm shift a1) (VarPTm var_zero))) ->
     R a0 a1 ->
     R a0 b.
   Proof. move => ->; apply AppEta. Qed.
 
-  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+  Lemma renaming (a b : PTm) (ξ : nat -> nat) :
     R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
-    move => h. move : m ξ.
-    elim : n a b /h.
+    move => h. move : ξ.
+    elim : a b /h.
     move => *; apply : AppAbs'; eauto; by asimpl.
     all : match goal with
           | [ |- context[var_zero]] => move => *; apply : AppEta'; eauto; by asimpl
@@ -109,23 +109,23 @@ Module Par.
   Qed.
 
 
-  Lemma morphing n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma morphing (a b : PTm) (ρ0 ρ1 : nat -> PTm) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
     R a b -> R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
   Proof.
-    move => + h. move : m ρ0 ρ1. elim : n a b/h.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ρ0 ρ1 hρ /=.
+    move => + h. move : ρ0 ρ1. elim : a b/h.
+    - move => a0 a1 b0 b1 ha iha hb ihb ρ0 ρ1 hρ /=.
       eapply AppAbs' with (a1 := subst_PTm (up_PTm_PTm ρ1) a1); eauto.
       by asimpl.
-      hauto l:on use:renaming inv:option.
+      hauto l:on use:renaming inv:nat.
     - hauto lq:on rew:off ctrs:R.
-    - hauto l:on inv:option use:renaming ctrs:R.
+    - hauto l:on inv:nat use:renaming ctrs:R.
     - hauto lq:on use:ProjPair'.
-    - move => n a0 a1 ha iha m ρ0 ρ1 hρ /=.
+    - move => a0 a1 ha iha ρ0 ρ1 hρ /=.
       apply : AppEta'; eauto. by asimpl.
     - hauto lq:on ctrs:R.
     - sfirstorder.
-    - hauto l:on inv:option ctrs:R use:renaming.
+    - hauto l:on inv:nat ctrs:R use:renaming.
     - hauto q:on ctrs:R.
     - qauto l:on ctrs:R.
     - qauto l:on ctrs:R.
@@ -134,16 +134,16 @@ Module Par.
     - qauto l:on ctrs:R.
   Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
     R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof. hauto l:on use:morphing, refl. Qed.
 
-  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
+  Lemma antirenaming (a : PTm) (b : PTm) (ξ : nat -> nat) :
     R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
   Proof.
     move E : (ren_PTm ξ a) => u h.
-    move : n ξ a E. elim : m u b/h.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ξ []//=.
+    move : ξ a E. elim : u b/h.
+    - move => a0 a1 b0 b1 ha iha hb ihb ξ []//=.
       move => c c0 [+ ?]. subst.
       case : c => //=.
       move => c [?]. subst.
@@ -153,7 +153,7 @@ Module Par.
       eexists. split.
       apply AppAbs; eauto.
       by asimpl.
-    - move => n a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc m ξ []//=.
+    - move => a0 a1 b0 b1 c0 c1 ha iha hb ihb hc ihc ξ []//=.
       move => []//= t t0 t1 [*]. subst.
       spec_refl.
       move : iha => [? [*]].
@@ -162,43 +162,43 @@ Module Par.
       eexists. split.
       apply AppPair; hauto. subst.
       by asimpl.
-    - move => n p a0 a1 ha iha m ξ []//= p0 []//= t [*]. subst.
+    - move => p a0 a1 ha iha ξ []//= p0 []//= t [*]. subst.
       spec_refl. move : iha => [b0 [? ?]]. subst.
       eexists. split. apply ProjAbs; eauto. by asimpl.
-    - move => n p a0 a1 b0 b1 ha iha hb ihb m ξ []//= p0 []//= t t0[*].
+    - move => p a0 a1 b0 b1 ha iha hb ihb ξ []//= p0 []//= t t0[*].
       subst. spec_refl.
       move : iha => [b0 [? ?]].
       move : ihb => [c0 [? ?]]. subst.
       eexists. split. by eauto using ProjPair.
       hauto q:on.
-    - move => n a0 a1 ha iha m ξ a ?. subst.
+    - move => a0 a1 ha iha ξ a ?. subst.
       spec_refl. move : iha => [a0 [? ?]]. subst.
       eexists. split. apply AppEta; eauto.
       by asimpl.
-    - move => n a0 a1 ha iha m ξ a ?. subst.
+    - move => a0 a1 ha iha ξ a ?. subst.
       spec_refl. move : iha => [b0 [? ?]]. subst.
       eexists. split. apply PairEta; eauto.
       by asimpl.
-    - move => n i m ξ []//=.
+    - move => i ξ []//=.
       hauto l:on.
-    - move => n a0 a1 ha iha m ξ []//= t [*]. subst.
+    - move => a0 a1 ha iha ξ []//= t [*]. subst.
       spec_refl.
       move  :iha => [b0 [? ?]]. subst.
       eexists. split. by apply AbsCong; eauto.
       done.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ξ []//= t t0 [*]. subst.
+    - move => a0 a1 b0 b1 ha iha hb ihb ξ []//= t t0 [*]. subst.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       move : ihb => [c0 [? ?]]. subst.
       eexists. split. by apply AppCong; eauto.
       done.
-    - move => n a0 a1 b0 b1 ha iha hb ihb m ξ []//= t t0[*]. subst.
+    - move => a0 a1 b0 b1 ha iha hb ihb ξ []//= t t0[*]. subst.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       move : ihb => [c0 [? ?]]. subst.
       eexists. split=>/=. by apply PairCong; eauto.
       done.
-    - move => n p a0 a1 ha iha m ξ []//= p0 t [*]. subst.
+    - move => p a0 a1 ha iha ξ []//= p0 t [*]. subst.
       spec_refl.
       move : iha => [b0 [? ?]]. subst.
       eexists. split. by apply ProjCong; eauto.
@@ -359,7 +359,7 @@ Module RPar.
     R a b ->
     (forall i, R (ρ0 i) (ρ1 i)) ->
     (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)).
-  Proof. hauto q:on inv:option. Qed.
+  Proof. hauto q:on inv:nat. Qed.
 
   Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
@@ -399,7 +399,7 @@ Module RPar.
     R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons d VarPTm) b).
   Proof.
     move => h0 h1. apply morphing => //=.
-    qauto l:on ctrs:R inv:option.
+    qauto l:on ctrs:R inv:nat.
   Qed.
 
   Lemma var_or_const_imp {n} (a b : PTm n) :
@@ -595,7 +595,7 @@ Module RPar'.
     R a b ->
     (forall i, R (ρ0 i) (ρ1 i)) ->
     (forall i, R ((scons a ρ0) i) ((scons b ρ1) i)).
-  Proof. hauto q:on inv:option. Qed.
+  Proof. hauto q:on inv:nat. Qed.
 
   Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
@@ -633,7 +633,7 @@ Module RPar'.
     R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons d VarPTm) b).
   Proof.
     move => h0 h1. apply morphing => //=.
-    qauto l:on ctrs:R inv:option.
+    qauto l:on ctrs:R inv:nat.
   Qed.
 
   Lemma var_or_const_imp {n} (a b : PTm n) :
@@ -786,7 +786,7 @@ Module ERed.
     move => h. move : m ρ. elim : n a b / h => n.
     move => a m ρ /=.
     apply : AppEta'; eauto. by asimpl.
-    all : hauto ctrs:R inv:option use:renaming.
+    all : hauto ctrs:R inv:nat use:renaming.
   Qed.
 
 End ERed.
@@ -892,11 +892,11 @@ Module EPar.
       apply : AppEta'; eauto. by asimpl.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
-    - hauto l:on ctrs:R use:renaming inv:option.
+    - hauto l:on ctrs:R use:renaming inv:nat.
     - hauto q:on ctrs:R.
     - hauto q:on ctrs:R.
     - hauto q:on ctrs:R.
-    - hauto l:on ctrs:R use:renaming inv:option.
+    - hauto l:on ctrs:R use:renaming inv:nat.
     - hauto lq:on ctrs:R.
     - hauto lq:on ctrs:R.
   Qed.
@@ -907,7 +907,7 @@ Module EPar.
     R (subst_PTm (scons b0 VarPTm) a0) (subst_PTm (scons b1 VarPTm) a1).
   Proof.
     move => h0 h1. apply morphing => //.
-    hauto lq:on ctrs:R inv:option.
+    hauto lq:on ctrs:R inv:nat.
   Qed.
 
 End EPar.
@@ -1018,7 +1018,7 @@ Module RPars.
   Lemma substing n (a b : PTm (S n)) c :
     rtc RPar.R a b ->
     rtc RPar.R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons c VarPTm) b).
-  Proof. hauto lq:on use:morphing inv:option. Qed.
+  Proof. hauto lq:on use:morphing inv:nat. Qed.
 
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
@@ -1099,7 +1099,7 @@ Module RPars'.
   Lemma substing n (a b : PTm (S n)) c :
     rtc RPar'.R a b ->
     rtc RPar'.R (subst_PTm (scons c VarPTm) a) (subst_PTm (scons c VarPTm) b).
-  Proof. hauto lq:on use:morphing inv:option. Qed.
+  Proof. hauto lq:on use:morphing inv:nat. Qed.
 
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ρ : fin n -> PTm m) :
     (forall i, var_or_const (ρ i)) ->
@@ -2328,7 +2328,7 @@ Proof.
       have : wn (subst_PTm (scons (PBot) VarPTm) a3) by sfirstorder.
       move => h. apply wn_abs.
       move : h. apply wn_antirenaming.
-      hauto lq:on rew:off inv:option.
+      hauto lq:on rew:off inv:nat.
     + hauto q:on inv:RPar'.R ctrs:rtc b:on.
 Qed.
 
diff --git a/theories/typing.v b/theories/typing.v
new file mode 100644
index 0000000..d41facd
--- /dev/null
+++ b/theories/typing.v
@@ -0,0 +1,251 @@
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+
+Reserved Notation "Γ ⊢ a ∈ A" (at level 70).
+Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70).
+Reserved Notation "Γ ⊢ A ≲ B" (at level 70).
+Reserved Notation "⊢ Γ" (at level 70).
+Inductive Wt : list PTm -> PTm -> PTm -> Prop :=
+| T_Var i Γ A :
+  ⊢ Γ ->
+  lookup i Γ A ->
+  Γ ⊢ VarPTm i ∈ A
+
+| T_Bind Γ i p (A : PTm) (B : PTm) :
+  Γ ⊢ A ∈ PUniv i ->
+  cons A Γ ⊢ B ∈ PUniv i ->
+  Γ ⊢ PBind p A B ∈ PUniv i
+
+| T_Abs Γ (a : PTm) A B i :
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  (cons A Γ) ⊢ a ∈ B ->
+  Γ ⊢ PAbs a ∈ PBind PPi A B
+
+| T_App Γ (b a : PTm) A B :
+  Γ ⊢ b ∈ PBind PPi A B ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ PApp b a ∈ subst_PTm (scons a VarPTm) B
+
+| T_Pair Γ (a b : PTm) A B i :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ PPair a b ∈ PBind PSig A B
+
+| T_Proj1 Γ (a : PTm) A B :
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ PProj PL a ∈ A
+
+| T_Proj2 Γ (a : PTm) A B :
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B
+
+| T_Univ Γ i :
+  ⊢ Γ ->
+  Γ ⊢ PUniv i ∈ PUniv (S i)
+
+| T_Nat Γ i :
+  ⊢ Γ ->
+  Γ ⊢ PNat ∈ PUniv i
+
+| T_Zero Γ :
+  ⊢ Γ ->
+  Γ ⊢ PZero ∈ PNat
+
+| T_Suc Γ (a : PTm) :
+  Γ ⊢ a ∈ PNat ->
+  Γ ⊢ PSuc a ∈ PNat
+
+| T_Ind Γ P (a : PTm) b c i :
+  cons PNat Γ ⊢ P ∈ PUniv i ->
+  Γ ⊢ a ∈ PNat ->
+  Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
+  (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) ->
+  Γ ⊢ PInd P a b c ∈ subst_PTm (scons a VarPTm) P
+
+| T_Conv Γ (a : PTm) A B :
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ A ≲ B ->
+  Γ ⊢ a ∈ B
+
+with Eq : list PTm -> PTm -> PTm -> PTm -> Prop :=
+(* Structural *)
+| E_Refl Γ (a : PTm ) A :
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ a ≡ a ∈ A
+
+| E_Symmetric Γ (a b : PTm) A :
+  Γ ⊢ a ≡ b ∈ A ->
+  Γ ⊢ b ≡ a ∈ A
+
+| E_Transitive Γ (a b c : PTm) A :
+  Γ ⊢ a ≡ b ∈ A ->
+  Γ ⊢ b ≡ c ∈ A ->
+  Γ ⊢ a ≡ c ∈ A
+
+(* Congruence *)
+| E_Bind Γ i p (A0 A1 : PTm) B0 B1 :
+  Γ ⊢ A0 ∈ PUniv i ->
+  Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
+  (cons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
+  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i
+
+| E_Abs Γ (a b : PTm) A B i :
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  (cons A Γ) ⊢ a ≡ b ∈ B ->
+  Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B
+
+| E_App Γ i (b0 b1 a0 a1 : PTm) A B :
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B ->
+  Γ ⊢ a0 ≡ a1 ∈ A ->
+  Γ ⊢ PApp b0 a0 ≡ PApp b1 a1 ∈ subst_PTm (scons a0 VarPTm) B
+
+| E_Pair Γ (a0 a1 b0 b1 : PTm) A B i :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a0 ≡ a1 ∈ A ->
+  Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons a0 VarPTm) B ->
+  Γ ⊢ PPair a0 b0 ≡ PPair a1 b1 ∈ PBind PSig A B
+
+| E_Proj1 Γ (a b : PTm) A B :
+  Γ ⊢ a ≡ b ∈ PBind PSig A B ->
+  Γ ⊢ PProj PL a ≡ PProj PL b ∈ A
+
+| E_Proj2 Γ i (a b : PTm) A B :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ≡ b ∈ PBind PSig A B ->
+  Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B
+
+| E_IndCong Γ P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 i :
+  (cons PNat Γ) ⊢ P0 ∈ PUniv i ->
+  (cons PNat Γ) ⊢ P0 ≡ P1 ∈ PUniv i ->
+  Γ ⊢ a0 ≡ a1 ∈ PNat ->
+  Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons PZero VarPTm) P0 ->
+  (cons P0 ((cons PNat Γ))) ⊢ c0 ≡ c1 ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P0) ->
+  Γ ⊢ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0
+
+| E_SucCong Γ (a b : PTm) :
+  Γ ⊢ a ≡ b ∈ PNat ->
+  Γ ⊢ PSuc a ≡ PSuc b ∈ PNat
+
+| E_Conv Γ (a b : PTm) A B :
+  Γ ⊢ a ≡ b ∈ A ->
+  Γ ⊢ A ≲ B ->
+  Γ ⊢ a ≡ b ∈ B
+
+(* Beta *)
+| E_AppAbs Γ (a : PTm) b A B i:
+  Γ ⊢ PBind PPi A B ∈ PUniv i ->
+  Γ ⊢ b ∈ A ->
+  (cons A Γ) ⊢ a ∈ B ->
+  Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ subst_PTm (scons b VarPTm ) B
+
+| E_ProjPair1 Γ (a b : PTm) A B i :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A
+
+| E_ProjPair2 Γ (a b : PTm) A B i :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ PProj PR (PPair a b) ≡ b ∈ subst_PTm (scons a VarPTm) B
+
+| E_IndZero Γ P i (b : PTm) c :
+  (cons PNat Γ) ⊢ P ∈ PUniv i ->
+  Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
+  (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) ->
+  Γ ⊢ PInd P PZero b c ≡ b ∈ subst_PTm (scons PZero VarPTm) P
+
+| E_IndSuc Γ P (a : PTm) b c i :
+  (cons PNat Γ) ⊢ P ∈ PUniv i ->
+  Γ ⊢ a ∈ PNat ->
+  Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
+  (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) ->
+  Γ ⊢ PInd P (PSuc a) b c ≡ (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ∈ subst_PTm (scons (PSuc a) VarPTm) P
+
+(* Eta *)
+| E_AppEta Γ (b : PTm) A B i :
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  Γ ⊢ b ∈ PBind PPi A B ->
+  Γ ⊢ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B
+
+| E_PairEta Γ (a : PTm ) A B i :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ a ≡ PPair (PProj PL a) (PProj PR a) ∈ PBind PSig A B
+
+with LEq : list PTm -> PTm -> PTm -> Prop :=
+(* Structural *)
+| Su_Transitive Γ (A B C : PTm) :
+  Γ ⊢ A ≲ B ->
+  Γ ⊢ B ≲ C ->
+  Γ ⊢ A ≲ C
+
+(* Congruence *)
+| Su_Univ Γ i j :
+  ⊢ Γ ->
+  i <= j ->
+  Γ ⊢ PUniv i ≲ PUniv j
+
+| Su_Pi Γ (A0 A1 : PTm) B0 B1 i :
+  Γ ⊢ A0 ∈ PUniv i ->
+  Γ ⊢ A1 ≲ A0 ->
+  (cons A0 Γ) ⊢ B0 ≲ B1 ->
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1
+
+| Su_Sig Γ (A0 A1 : PTm) B0 B1 i :
+  Γ ⊢ A1 ∈ PUniv i ->
+  Γ ⊢ A0 ≲ A1 ->
+  (cons A1 Γ) ⊢ B0 ≲ B1 ->
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1
+
+(* Injecting from equalities *)
+| Su_Eq Γ (A : PTm) B i  :
+  Γ ⊢ A ≡ B ∈ PUniv i ->
+  Γ ⊢ A ≲ B
+
+(* Projection axioms *)
+| Su_Pi_Proj1 Γ (A0 A1 : PTm) B0 B1 :
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
+  Γ ⊢ A1 ≲ A0
+
+| Su_Sig_Proj1 Γ (A0 A1 : PTm) B0 B1 :
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
+  Γ ⊢ A0 ≲ A1
+
+| Su_Pi_Proj2 Γ (a0 a1 A0 A1 : PTm ) B0 B1 :
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
+  Γ ⊢ a0 ≡ a1 ∈ A1 ->
+  Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≲ subst_PTm (scons a1 VarPTm) B1
+
+| Su_Sig_Proj2 Γ (a0 a1 A0 A1 : PTm) B0 B1 :
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
+  Γ ⊢ a0 ≡ a1 ∈ A0 ->
+  Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≲ subst_PTm (scons a1 VarPTm) B1
+
+with Wff : list PTm -> Prop :=
+| Wff_Nil :
+  ⊢ nil
+| Wff_Cons Γ (A : PTm) i :
+  ⊢ Γ ->
+  Γ ⊢ A ∈ PUniv i ->
+  (* -------------------------------- *)
+  ⊢ (cons A Γ)
+
+where
+"Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A) and "Γ ⊢ A ≲ B" := (LEq Γ A B).
+
+Scheme wf_ind := Induction for Wff Sort Prop
+  with wt_ind := Induction for Wt Sort Prop
+  with eq_ind := Induction for Eq Sort Prop
+  with le_ind := Induction for LEq Sort Prop.
+
+Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind, le_ind.
+
+(* Lemma lem : *)
+(*   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> ...) /\ *)
+(*   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> ...)  /\ *)
+(*   (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> ...) /\ *)
+(*   (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> ...). *)
+(* Proof. apply wt_mutual. ... *)