From 657c1325c99aa5891ba440af7ede8e3f656068c6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 2 Mar 2025 16:40:39 -0500
Subject: [PATCH 01/24] Add unscoped syntax

---
 Makefile                       |   4 +-
 syntax.sig                     |   1 -
 theories/Autosubst2/syntax.v   | 785 ++++++++++++---------------------
 theories/Autosubst2/unscoped.v | 213 +++++++++
 theories/common.v              |  70 +--
 5 files changed, 532 insertions(+), 541 deletions(-)
 create mode 100644 theories/Autosubst2/unscoped.v

diff --git a/Makefile b/Makefile
index ef5b76f..79b3720 100644
--- a/Makefile
+++ b/Makefile
@@ -16,12 +16,12 @@ 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
+	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/unscoped.v
 
 FORCE:
diff --git a/syntax.sig b/syntax.sig
index a50911e..24931eb 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -16,7 +16,6 @@ PPair : PTm -> PTm -> PTm
 PProj : PTag -> PTm -> PTm
 PBind : BTag -> PTm -> (bind PTm in PTm) -> PTm
 PUniv : nat -> PTm
-PBot : PTm
 PNat : PTm
 PZero : PTm
 PSuc : PTm -> PTm
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index 5d04c05..ee8b076 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,221 +33,168 @@ 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
-  | PBind : BTag -> PTm n_PTm -> PTm (S n_PTm) -> PTm n_PTm
-  | PUniv : nat -> PTm n_PTm
-  | PBot : PTm n_PTm
-  | PNat : PTm n_PTm
-  | PZero : PTm n_PTm
-  | PSuc : PTm n_PTm -> PTm n_PTm
-  | PInd :
-      PTm (S n_PTm) ->
-      PTm n_PTm -> PTm n_PTm -> PTm (S (S n_PTm)) -> PTm n_PTm.
+Inductive PTm : Type :=
+  | VarPTm : nat -> PTm
+  | PAbs : PTm -> PTm
+  | PApp : PTm -> PTm -> PTm
+  | PPair : PTm -> PTm -> PTm
+  | PProj : PTag -> PTm -> PTm
+  | PBind : BTag -> PTm -> PTm -> PTm
+  | PUniv : nat -> PTm
+  | PNat : PTm
+  | PZero : PTm
+  | PSuc : PTm -> PTm
+  | PInd : PTm -> PTm -> PTm -> PTm -> 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_PBind {m_PTm : nat} {s0 : BTag} {s1 : PTm m_PTm}
-  {s2 : PTm (S m_PTm)} {t0 : BTag} {t1 : PTm m_PTm} {t2 : PTm (S m_PTm)}
-  (H0 : s0 = t0) (H1 : s1 = t1) (H2 : s2 = t2) :
-  PBind m_PTm s0 s1 s2 = PBind m_PTm t0 t1 t2.
+Lemma congr_PBind {s0 : BTag} {s1 : PTm} {s2 : PTm} {t0 : BTag} {t1 : PTm}
+  {t2 : PTm} (H0 : s0 = t0) (H1 : s1 = t1) (H2 : s2 = t2) :
+  PBind s0 s1 s2 = PBind t0 t1 t2.
 Proof.
 exact (eq_trans
-         (eq_trans (eq_trans eq_refl (ap (fun x => PBind m_PTm x s1 s2) H0))
-            (ap (fun x => PBind m_PTm t0 x s2) H1))
-         (ap (fun x => PBind m_PTm t0 t1 x) H2)).
+         (eq_trans (eq_trans eq_refl (ap (fun x => PBind x s1 s2) H0))
+            (ap (fun x => PBind t0 x s2) H1))
+         (ap (fun x => PBind t0 t1 x) H2)).
 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_PNat : PNat = PNat.
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma congr_PNat {m_PTm : nat} : PNat m_PTm = PNat m_PTm.
+Lemma congr_PZero : PZero = PZero.
 Proof.
 exact (eq_refl).
 Qed.
 
-Lemma congr_PZero {m_PTm : nat} : PZero m_PTm = PZero m_PTm.
+Lemma congr_PSuc {s0 : PTm} {t0 : PTm} (H0 : s0 = t0) : PSuc s0 = PSuc t0.
 Proof.
-exact (eq_refl).
+exact (eq_trans eq_refl (ap (fun x => PSuc x) H0)).
 Qed.
 
-Lemma congr_PSuc {m_PTm : nat} {s0 : PTm m_PTm} {t0 : PTm m_PTm}
-  (H0 : s0 = t0) : PSuc m_PTm s0 = PSuc m_PTm t0.
-Proof.
-exact (eq_trans eq_refl (ap (fun x => PSuc m_PTm x) H0)).
-Qed.
-
-Lemma congr_PInd {m_PTm : nat} {s0 : PTm (S m_PTm)} {s1 : PTm m_PTm}
-  {s2 : PTm m_PTm} {s3 : PTm (S (S m_PTm))} {t0 : PTm (S m_PTm)}
-  {t1 : PTm m_PTm} {t2 : PTm m_PTm} {t3 : PTm (S (S m_PTm))} (H0 : s0 = t0)
-  (H1 : s1 = t1) (H2 : s2 = t2) (H3 : s3 = t3) :
-  PInd m_PTm s0 s1 s2 s3 = PInd m_PTm t0 t1 t2 t3.
+Lemma congr_PInd {s0 : PTm} {s1 : PTm} {s2 : PTm} {s3 : PTm} {t0 : PTm}
+  {t1 : PTm} {t2 : PTm} {t3 : PTm} (H0 : s0 = t0) (H1 : s1 = t1)
+  (H2 : s2 = t2) (H3 : s3 = t3) : PInd s0 s1 s2 s3 = PInd t0 t1 t2 t3.
 Proof.
 exact (eq_trans
          (eq_trans
-            (eq_trans
-               (eq_trans eq_refl (ap (fun x => PInd m_PTm x s1 s2 s3) H0))
-               (ap (fun x => PInd m_PTm t0 x s2 s3) H1))
-            (ap (fun x => PInd m_PTm t0 t1 x s3) H2))
-         (ap (fun x => PInd m_PTm t0 t1 t2 x) H3)).
+            (eq_trans (eq_trans eq_refl (ap (fun x => PInd x s1 s2 s3) H0))
+               (ap (fun x => PInd t0 x s2 s3) H1))
+            (ap (fun x => PInd t0 t1 x s3) H2))
+         (ap (fun x => PInd t0 t1 t2 x) H3)).
 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)
-  | PBind _ s0 s1 s2 =>
-      PBind n_PTm s0 (ren_PTm xi_PTm s1) (ren_PTm (upRen_PTm_PTm xi_PTm) s2)
-  | PUniv _ s0 => PUniv n_PTm s0
-  | PBot _ => PBot n_PTm
-  | PNat _ => PNat n_PTm
-  | PZero _ => PZero n_PTm
-  | PSuc _ s0 => PSuc n_PTm (ren_PTm xi_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
-      PInd n_PTm (ren_PTm (upRen_PTm_PTm xi_PTm) s0) (ren_PTm xi_PTm s1)
+  | 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)
+  | PBind s0 s1 s2 =>
+      PBind s0 (ren_PTm xi_PTm s1) (ren_PTm (upRen_PTm_PTm xi_PTm) s2)
+  | PUniv s0 => PUniv s0
+  | PNat => PNat
+  | PZero => PZero
+  | PSuc s0 => PSuc (ren_PTm xi_PTm s0)
+  | PInd s0 s1 s2 s3 =>
+      PInd (ren_PTm (upRen_PTm_PTm xi_PTm) s0) (ren_PTm xi_PTm s1)
         (ren_PTm xi_PTm s2)
         (ren_PTm (upRen_PTm_PTm (upRen_PTm_PTm xi_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
-      PBind n_PTm s0 (subst_PTm sigma_PTm s1)
-        (subst_PTm (up_PTm_PTm sigma_PTm) s2)
-  | PUniv _ s0 => PUniv n_PTm s0
-  | PBot _ => PBot n_PTm
-  | PNat _ => PNat n_PTm
-  | PZero _ => PZero n_PTm
-  | PSuc _ s0 => PSuc n_PTm (subst_PTm sigma_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
-      PInd n_PTm (subst_PTm (up_PTm_PTm sigma_PTm) s0)
-        (subst_PTm sigma_PTm s1) (subst_PTm sigma_PTm s2)
+  | 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)
+  | PBind s0 s1 s2 =>
+      PBind s0 (subst_PTm sigma_PTm s1) (subst_PTm (up_PTm_PTm sigma_PTm) s2)
+  | PUniv s0 => PUniv s0
+  | PNat => PNat
+  | PZero => PZero
+  | PSuc s0 => PSuc (subst_PTm sigma_PTm s0)
+  | PInd s0 s1 s2 s3 =>
+      PInd (subst_PTm (up_PTm_PTm sigma_PTm) s0) (subst_PTm sigma_PTm s1)
+        (subst_PTm sigma_PTm s2)
         (subst_PTm (up_PTm_PTm (up_PTm_PTm sigma_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
+  | PProj s0 s1 => congr_PProj (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1)
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1)
         (idSubst_PTm (up_PTm_PTm sigma_PTm) (upId_PTm_PTm _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 => congr_PSuc (idSubst_PTm sigma_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 => congr_PSuc (idSubst_PTm sigma_PTm Eq_PTm s0)
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (idSubst_PTm (up_PTm_PTm sigma_PTm) (upId_PTm_PTm _ Eq_PTm) s0)
         (idSubst_PTm sigma_PTm Eq_PTm s1) (idSubst_PTm sigma_PTm Eq_PTm s2)
@@ -255,54 +202,42 @@ Fixpoint idSubst_PTm {m_PTm : nat} (sigma_PTm : fin m_PTm -> PTm m_PTm)
            (upId_PTm_PTm _ (upId_PTm_PTm _ Eq_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0) (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
         (extRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
            (upExtRen_PTm_PTm _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 => congr_PSuc (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 => congr_PSuc (extRen_PTm xi_PTm zeta_PTm Eq_PTm s0)
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (extRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
            (upExtRen_PTm_PTm _ _ Eq_PTm) s0)
@@ -313,55 +248,43 @@ ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
            (upExtRen_PTm_PTm _ _ (upExtRen_PTm_PTm _ _ Eq_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0) (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
         (ext_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
            (upExt_PTm_PTm _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 => congr_PSuc (ext_PTm sigma_PTm tau_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 => congr_PSuc (ext_PTm sigma_PTm tau_PTm Eq_PTm s0)
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (ext_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
            (upExt_PTm_PTm _ _ Eq_PTm) s0)
@@ -372,57 +295,43 @@ subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
            (upExt_PTm_PTm _ _ (upExt_PTm_PTm _ _ Eq_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0)
         (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
         (compRenRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
            (upRen_PTm_PTm rho_PTm) (up_ren_ren _ _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 =>
-      congr_PSuc (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 => congr_PSuc (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s0)
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (compRenRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
            (upRen_PTm_PTm rho_PTm) (up_ren_ren _ _ _ Eq_PTm) s0)
@@ -434,66 +343,48 @@ Fixpoint compRenRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
            (up_ren_ren _ _ _ (up_ren_ren _ _ _ Eq_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0)
         (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
         (compRenSubst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm tau_PTm)
            (up_PTm_PTm theta_PTm) (up_ren_subst_PTm_PTm _ _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 =>
       congr_PSuc (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (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)
@@ -506,9 +397,8 @@ Fixpoint compRenSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
            s3)
   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 =
@@ -516,76 +406,50 @@ 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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0)
         (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
         (compSubstRen_PTm (up_PTm_PTm sigma_PTm) (upRen_PTm_PTm zeta_PTm)
            (up_PTm_PTm theta_PTm) (up_subst_ren_PTm_PTm _ _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 =>
       congr_PSuc (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (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)
@@ -598,9 +462,8 @@ ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
            s3)
   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 =
@@ -608,78 +471,51 @@ 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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0)
         (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
         (compSubstSubst_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
            (up_PTm_PTm theta_PTm) (up_subst_subst_PTm_PTm _ _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 =>
       congr_PSuc (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (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)
@@ -692,126 +528,101 @@ subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
               (up_subst_subst_PTm_PTm _ _ _ Eq_PTm)) s3)
   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)
-  | PBind _ s0 s1 s2 =>
+  | PBind s0 s1 s2 =>
       congr_PBind (eq_refl s0) (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
         (rinst_inst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm sigma_PTm)
            (rinstInst_up_PTm_PTm _ _ Eq_PTm) s2)
-  | PUniv _ s0 => congr_PUniv (eq_refl s0)
-  | PBot _ => congr_PBot
-  | PNat _ => congr_PNat
-  | PZero _ => congr_PZero
-  | PSuc _ s0 => congr_PSuc (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0)
-  | PInd _ s0 s1 s2 s3 =>
+  | PUniv s0 => congr_PUniv (eq_refl s0)
+  | PNat => congr_PNat
+  | PZero => congr_PZero
+  | PSuc s0 => congr_PSuc (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s0)
+  | PInd s0 s1 s2 s3 =>
       congr_PInd
         (rinst_inst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm sigma_PTm)
            (rinstInst_up_PTm_PTm _ _ Eq_PTm) s0)
@@ -822,71 +633,61 @@ Fixpoint rinst_inst_PTm {m_PTm : nat} {n_PTm : nat}
            (rinstInst_up_PTm_PTm _ _ (rinstInst_up_PTm_PTm _ _ Eq_PTm)) s3)
   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.
@@ -894,18 +695,14 @@ Qed.
 Class Up_PTm X Y :=
     up_PTm : X -> Y.
 
-#[global]
-Instance Subst_PTm  {m_PTm n_PTm : nat}: (Subst1 _ _ _) :=
- (@subst_PTm m_PTm n_PTm).
+#[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_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_PTm ]" := (subst_PTm sigma_PTm)
 ( at level 1, left associativity, only printing)  : fscope.
@@ -932,9 +729,9 @@ Notation "x '__PTm'" := (VarPTm x) ( at level 5, format "x __PTm")  :
 subst_scope.
 
 #[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')
@@ -942,17 +739,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')
@@ -960,9 +756,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.
@@ -994,9 +790,7 @@ Ltac asimpl' := repeat (first
                  | progress setoid_rewrite rinstId'_PTm
                  | progress setoid_rewrite instId'_PTm_pointwise
                  | progress setoid_rewrite instId'_PTm
-                 | progress
-                    unfold up_list_PTm_PTm, up_PTm_PTm, upRen_list_PTm_PTm,
-                     upRen_PTm_PTm, up_ren
+                 | progress unfold up_PTm_PTm, upRen_PTm_PTm, up_ren
                  | progress cbn[subst_PTm ren_PTm]
                  | progress fsimpl ]).
 
@@ -1021,36 +815,11 @@ End Core.
 
 Module Extra.
 
-Import
-Core.
+Import Core.
 
-Arguments VarPTm {n_PTm}.
+#[global] Hint Opaque subst_PTm: rewrite.
 
-Arguments PInd {n_PTm}.
-
-Arguments PSuc {n_PTm}.
-
-Arguments PZero {n_PTm}.
-
-Arguments PNat {n_PTm}.
-
-Arguments PBot {n_PTm}.
-
-Arguments PUniv {n_PTm}.
-
-Arguments PBind {n_PTm}.
-
-Arguments PProj {n_PTm}.
-
-Arguments PPair {n_PTm}.
-
-Arguments PApp {n_PTm}.
-
-Arguments PAbs {n_PTm}.
-
-#[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/common.v b/theories/common.v
index 282038d..79b0b19 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -1,41 +1,52 @@
-Require Import Autosubst2.fintype Autosubst2.syntax Autosubst2.core ssreflect.
+Require Import Autosubst2.unscoped Autosubst2.syntax Autosubst2.core ssreflect.
 From Ltac2 Require Ltac2.
 Import Ltac2.Notations.
 Import Ltac2.Control.
 From Hammer Require Import Tactics.
 
-Definition renaming_ok {n m} (Γ : fin n -> PTm n) (Δ : fin m -> PTm m) (ξ : fin m -> fin n) :=
-  forall (i : fin m), ren_PTm ξ (Δ i) = Γ (ξ i).
+Inductive lookup : nat -> list PTm -> PTm -> Prop :=
+| here A Γ : lookup 0 (cons A Γ) (ren_PTm shift A)
+| there i Γ A B :
+  lookup i Γ A ->
+  lookup (S i) (cons B Γ) (ren_PTm shift A).
 
-Lemma up_injective n m (ξ : fin n -> fin m) :
-  (forall i j, ξ i = ξ j -> i = j) ->
-  forall i j, (upRen_PTm_PTm ξ)  i = (upRen_PTm_PTm ξ) j -> i = j.
+Definition renaming_ok (Γ : list PTm) (Δ : list PTm) (ξ : nat -> nat) :=
+  forall i A, lookup i Δ A -> lookup (ξ i) Γ (ren_PTm ξ A).
+
+Definition ren_inj (ξ : nat -> nat) := forall i j, ξ i = ξ j -> i = j.
+
+Lemma up_injective (ξ : nat -> nat) :
+  ren_inj ξ ->
+  ren_inj (upRen_PTm_PTm ξ).
 Proof.
-  sblast inv:option.
+  move => h i j.
+  case : i => //=; case : j => //=.
+  move => i j. rewrite /funcomp. hauto lq:on rew:off unfold:ren_inj.
 Qed.
 
 Local Ltac2 rec solve_anti_ren () :=
   let x := Fresh.in_goal (Option.get (Ident.of_string "x")) in
   intro $x;
   lazy_match! Constr.type (Control.hyp x) with
-  | fin _ -> _ _ => (ltac1:(case;hauto lq:on rew:off use:up_injective))
+  | nat -> nat => (ltac1:(case => *//=; qauto l:on use:up_injective unfold:ren_inj))
   | _ => solve_anti_ren ()
   end.
 
 Local Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
 
-Lemma ren_injective n m (a b : PTm n) (ξ : fin n -> fin m) :
-  (forall i j, ξ i = ξ j -> i = j) ->
+Lemma ren_injective (a b : PTm) (ξ : nat -> nat) :
+  ren_inj ξ ->
   ren_PTm ξ a = ren_PTm ξ b ->
   a = b.
 Proof.
-  move : m ξ b. elim : n / a => //; try solve_anti_ren.
+  move : ξ b. elim : a => //; try solve_anti_ren.
+  move => p ihp ξ []//=. hauto lq:on inv:PTm, nat ctrs:- use:up_injective.
 Qed.
 
 Inductive HF : Set :=
 | H_Pair | H_Abs | H_Univ | H_Bind (p : BTag) | H_Nat | H_Suc | H_Zero | H_Bot.
 
-Definition ishf {n} (a : PTm n) :=
+Definition ishf (a : PTm) :=
   match a with
   | PPair _ _ => true
   | PAbs _ => true
@@ -47,7 +58,7 @@ Definition ishf {n} (a : PTm n) :=
   | _ => false
   end.
 
-Definition toHF {n} (a : PTm n) :=
+Definition toHF (a : PTm) :=
   match a with
   | PPair _ _ => H_Pair
   | PAbs _ => H_Abs
@@ -59,54 +70,53 @@ Definition toHF {n} (a : PTm n) :=
   | _ => H_Bot
   end.
 
-Fixpoint ishne {n} (a : PTm n) :=
+Fixpoint ishne (a : PTm) :=
   match a with
   | VarPTm _ => true
   | PApp a _ => ishne a
   | PProj _ a => ishne a
-  | PBot => true
   | PInd _ n _ _ => ishne n
   | _ => false
   end.
 
-Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
+Definition isbind (a : PTm) := if a is PBind _ _ _ then true else false.
 
-Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
+Definition isuniv (a : PTm) := if a is PUniv _ then true else false.
 
-Definition ispair {n} (a : PTm n) :=
+Definition ispair (a : PTm) :=
   match a with
   | PPair _ _ => true
   | _ => false
   end.
 
-Definition isnat {n} (a : PTm n) := if a is PNat then true else false.
+Definition isnat (a : PTm) := if a is PNat then true else false.
 
-Definition iszero {n} (a : PTm n) := if a is PZero then true else false.
+Definition iszero (a : PTm) := if a is PZero then true else false.
 
-Definition issuc {n} (a : PTm n) := if a is PSuc _ then true else false.
+Definition issuc (a : PTm) := if a is PSuc _ then true else false.
 
-Definition isabs {n} (a : PTm n) :=
+Definition isabs (a : PTm) :=
   match a with
   | PAbs _ => true
   | _ => false
   end.
 
-Definition ishf_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+Definition ishf_ren (a : PTm)  (ξ : nat -> nat) :
   ishf (ren_PTm ξ a) = ishf a.
 Proof. case : a => //=. Qed.
 
-Definition isabs_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+Definition isabs_ren (a : PTm)  (ξ : nat -> nat) :
   isabs (ren_PTm ξ a) = isabs a.
 Proof. case : a => //=. Qed.
 
-Definition ispair_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+Definition ispair_ren (a : PTm)  (ξ : nat -> nat) :
   ispair (ren_PTm ξ a) = ispair a.
 Proof. case : a => //=. Qed.
 
-Definition ishne_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+Definition ishne_ren (a : PTm)  (ξ : nat -> nat) :
   ishne (ren_PTm ξ a) = ishne a.
-Proof. move : m ξ. elim : n / a => //=. Qed.
+Proof. move : ξ. elim : a => //=. Qed.
 
-Lemma renaming_shift n m Γ (ρ : fin n -> PTm m) A :
-  renaming_ok (funcomp (ren_PTm shift) (scons (subst_PTm ρ A) Γ)) Γ shift.
-Proof. sfirstorder. Qed.
+Lemma renaming_shift Γ (ρ : nat -> PTm) A :
+  renaming_ok (cons A Γ) Γ shift.
+Proof. rewrite /renaming_ok. hauto lq:on ctrs:lookup. Qed.

From 226b196d159bb45dd1e04a7ebaeadc46472669f7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 2 Mar 2025 17:35:51 -0500
Subject: [PATCH 02/24] Refactor half of fp_red

---
 theories/Autosubst2/fintype.v | 419 --------------
 theories/fp_red.v             | 988 +++++++++++++++++-----------------
 2 files changed, 486 insertions(+), 921 deletions(-)
 delete mode 100644 theories/Autosubst2/fintype.v

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/fp_red.v b/theories/fp_red.v
index bffe1a7..5f7ea0d 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -8,7 +8,7 @@ Require Import Arith.Wf_nat (well_founded_lt_compat).
 Require Import Psatz.
 From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn).
 From Hammer Require Import Tactics.
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 Require Import Btauto.
 Require Import Cdcl.Itauto.
 
@@ -23,7 +23,7 @@ Ltac2 spec_refl () :=
 Ltac spec_refl := ltac2:(Control.enter spec_refl).
 
 Module EPar.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (****************** Eta ***********************)
   | AppEta a0 a1 :
     R a0 a1 ->
@@ -54,8 +54,6 @@ Module EPar.
     R A0 A1 ->
     R B0 B1 ->
     R (PBind p A0 B0) (PBind p A1 B1)
-  | BotCong :
-    R PBot PBot
   | NatCong :
     R PNat PNat
   | IndCong P0 P1 a0 a1 b0 b1 c0 c1 :
@@ -72,64 +70,64 @@ Module EPar.
     (* ------------ *)
     R (PSuc a0) (PSuc a1).
 
-  Lemma refl n (a : PTm n) : R a a.
+  Lemma refl (a : PTm) : R a a.
   Proof.
-    elim : n / a; hauto lq:on ctrs:R.
+    elim : a; hauto lq:on ctrs:R.
   Qed.
 
-  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
+  Derive Dependent Inversion inv with (forall (a b : PTm), R a b) Sort Prop.
 
-  Lemma AppEta' n a0 a1 (u : PTm n) :
+  Lemma AppEta' a0 a1 (u : PTm) :
     u = (PAbs (PApp (ren_PTm shift a0) (VarPTm var_zero))) ->
     R a0 a1 ->
     R u a1.
   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.
 
     all : try qauto ctrs:R.
 
-    move => n a0 a1 ha iha m ξ /=.
+    move => a0 a1 ha iha ξ /=.
     eapply AppEta'; eauto. by asimpl.
   Qed.
 
-  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
+  Lemma morphing_ren (ρ0 ρ1 : nat -> PTm) (ξ : nat -> nat) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R ((funcomp (ren_PTm ξ) ρ0) i) ((funcomp (ren_PTm ξ) ρ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 -> PTm m) a b  :
+  Lemma morphing_ext (ρ0 ρ1 : nat -> PTm) 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.
+  Proof. hauto q:on inv:nat. Qed.
 
-  Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma morphing_up (ρ0 ρ1 : nat -> PTm) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
     (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 : 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 => n.
-    move => a0 a1 ha iha m ρ0 ρ1 hρ /=.
+    move => + h. move : ρ0 ρ1. elim : a b / h.
+    move => a0 a1 ha iha ρ0 ρ1 hρ /=.
     eapply AppEta'; eauto. by asimpl.
     all : hauto lq:on ctrs:R use:morphing_up.
   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.
 
 End EPar.
 
-Inductive SNe {n} : PTm n -> Prop :=
+Inductive SNe : PTm -> Prop :=
 | N_Var i :
   SNe (VarPTm i)
 | N_App a b :
@@ -139,15 +137,13 @@ Inductive SNe {n} : PTm n -> Prop :=
 | N_Proj p a :
   SNe a ->
   SNe (PProj p a)
-| N_Bot :
-  SNe PBot
 | N_Ind P a b c :
   SN P ->
   SNe a ->
   SN b ->
   SN c ->
   SNe (PInd P a b c)
-with SN  {n} : PTm n -> Prop :=
+with SN  : PTm -> Prop :=
 | N_Pair a b :
   SN a ->
   SN b ->
@@ -175,7 +171,7 @@ with SN  {n} : PTm n -> Prop :=
 | N_Suc a :
   SN a ->
   SN (PSuc a)
-with TRedSN {n} : PTm n -> PTm n -> Prop :=
+with TRedSN : PTm -> PTm -> Prop :=
 | N_β a b :
   SN b ->
   TRedSN (PApp (PAbs a) b) (subst_PTm (scons b VarPTm) a)
@@ -210,38 +206,38 @@ with TRedSN {n} : PTm n -> PTm n -> Prop :=
   TRedSN a0 a1 ->
   TRedSN (PInd P a0 b c) (PInd P a1 b c).
 
-Derive Dependent Inversion tred_inv with (forall n (a b : PTm n), TRedSN a b) Sort Prop.
+Derive Inversion tred_inv with (forall (a b : PTm), TRedSN a b) Sort Prop.
 
-Inductive SNat {n} : PTm n -> Prop :=
+Inductive SNat : PTm -> Prop :=
 | S_Zero : SNat PZero
 | S_Neu a : SNe a -> SNat a
 | S_Suc a : SNat a -> SNat (PSuc a)
 | S_Red a b : TRedSN a b -> SNat b -> SNat a.
 
-Lemma PProj_imp n p a :
-  @ishf n a ->
+Lemma PProj_imp p a :
+  ishf a ->
   ~~ ispair a ->
   ~ SN (PProj p a).
 Proof.
   move => + + h. move E  : (PProj p a) h => u h.
   move : p a E.
-  elim : n u / h => //=.
+  elim : u / h => //=.
   hauto lq:on inv:SNe,PTm.
   hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma PApp_imp n a b :
-  @ishf n a ->
+Lemma PApp_imp a b :
+  ishf a ->
   ~~ isabs a ->
   ~ SN (PApp a b).
 Proof.
   move => + + h. move E : (PApp a b) h => u h.
-  move : a b E. elim  : n u /h=>//=.
+  move : a b E. elim  : u /h=>//=.
   hauto lq:on inv:SNe,PTm.
   hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma PInd_imp n P (a : PTm n) b c :
+Lemma PInd_imp P (a : PTm) b c :
   ishf a ->
   ~~ iszero a ->
   ~~ issuc a ->
@@ -249,35 +245,35 @@ Lemma PInd_imp n P (a : PTm n) b c :
 Proof.
   move => + + + h. move E : (PInd P a b c) h => u h.
   move : P a b c E.
-  elim : n u /h => //=.
+  elim : u /h => //=.
   hauto lq:on inv:SNe,PTm.
   hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma PProjAbs_imp n p (a : PTm (S n)) :
+Lemma PProjAbs_imp p (a : PTm) :
   ~ SN (PProj p (PAbs a)).
 Proof.
   sfirstorder use:PProj_imp.
 Qed.
 
-Lemma PAppPair_imp n (a b0 b1 : PTm n ) :
+Lemma PAppPair_imp (a b0 b1 : PTm ) :
   ~ SN (PApp (PPair b0 b1) a).
 Proof.
   sfirstorder use:PApp_imp.
 Qed.
 
-Lemma PAppBind_imp n p (A : PTm n) B b :
+Lemma PAppBind_imp p (A : PTm) B b :
   ~ SN (PApp (PBind p A B) b).
 Proof.
   sfirstorder use:PApp_imp.
 Qed.
 
-Lemma PProjBind_imp n p p' (A : PTm n) B :
+Lemma PProjBind_imp p p' (A : PTm) B :
   ~ SN (PProj p (PBind p' A B)).
 Proof.
   move E :(PProj p (PBind p' A B)) => u hu.
   move : p p' A B E.
-  elim : n u /hu=>//=.
+  elim : u /hu=>//=.
   hauto lq:on inv:SNe.
   hauto lq:on inv:TRedSN.
 Qed.
@@ -288,7 +284,7 @@ Scheme sne_ind := Induction for SNe Sort Prop
 
 Combined Scheme sn_mutual from sne_ind, sn_ind, sred_ind.
 
-Fixpoint ne {n} (a : PTm n) :=
+Fixpoint ne (a : PTm) :=
   match a with
   | VarPTm i => true
   | PApp a b => ne a && nf b
@@ -297,13 +293,12 @@ Fixpoint ne {n} (a : PTm n) :=
   | PProj _ a => ne a
   | PUniv _ => false
   | PBind _ _ _ => false
-  | PBot => true
   | PInd P a b c => nf P && ne a && nf b && nf c
   | PNat => false
   | PSuc a => false
   | PZero => false
   end
-with nf {n} (a : PTm n) :=
+with nf (a : PTm) :=
   match a with
   | VarPTm i => true
   | PApp a b => ne a && nf b
@@ -312,69 +307,66 @@ with nf {n} (a : PTm n) :=
   | PProj _ a => ne a
   | PUniv _ => true
   | PBind _ A B => nf A && nf B
-  | PBot => true
   | PInd P a b c => nf P && ne a && nf b && nf c
   | PNat => true
   | PSuc a => nf a
   | PZero => true
 end.
 
-Lemma ne_nf n a : @ne n a -> nf a.
+Lemma ne_nf a : ne a -> nf a.
 Proof. elim : a => //=. Qed.
 
-Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
+Lemma ne_nf_ren (a : PTm) (ξ : nat -> nat) :
   (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)).
 Proof.
-  move : m ξ. elim : n / a => //=; solve [hauto b:on].
+  move : ξ. elim : a => //=; solve [hauto b:on].
 Qed.
 
-Inductive TRedSN' {n} (a : PTm n) : PTm n -> Prop :=
+Inductive TRedSN' (a : PTm) : PTm -> Prop :=
 | T_Refl :
   TRedSN' a a
 | T_Once b :
   TRedSN a b ->
   TRedSN' a b.
 
-Lemma SN_Proj n p (a : PTm n) :
+Lemma SN_Proj p (a : PTm) :
   SN (PProj p a) -> SN a.
 Proof.
   move E : (PProj p a) => u h.
   move : a E.
-  elim : n u / h => n //=; sauto.
+  elim : u / h => n //=; sauto.
 Qed.
 
-Lemma N_β' n a (b : PTm n) u :
+Lemma N_β' a (b : PTm) u :
   u = (subst_PTm (scons b VarPTm) a) ->
   SN b ->
   TRedSN (PApp (PAbs a) b) u.
 Proof. move => ->. apply N_β. Qed.
 
-Lemma N_IndSuc' n P a b c u :
+Lemma N_IndSuc' P a b c u :
   u = (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ->
   SN P ->
-  @SN n a ->
+  SN a ->
   SN b ->
   SN c ->
   TRedSN (PInd P (PSuc a) b c) u.
 Proof. move => ->. apply N_IndSuc. Qed.
 
-Lemma sn_renaming n :
-  (forall (a : PTm n) (s : SNe a), forall m (ξ : fin n -> fin m), SNe (ren_PTm ξ a)) /\
-  (forall (a : PTm n) (s : SN a), forall m (ξ : fin n -> fin m), SN (ren_PTm ξ a)) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall m (ξ : fin n -> fin m), TRedSN (ren_PTm ξ a) (ren_PTm ξ b)).
+Lemma sn_renaming :
+  (forall (a : PTm) (s : SNe a), forall (ξ : nat -> nat), SNe (ren_PTm ξ a)) /\
+  (forall (a : PTm) (s : SN a), forall (ξ : nat -> nat), SN (ren_PTm ξ a)) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall (ξ : nat -> nat), TRedSN (ren_PTm ξ a) (ren_PTm ξ b)).
 Proof.
-  move : n. apply sn_mutual => n; try qauto ctrs:SN, SNe, TRedSN depth:1.
-  move => a b ha iha m ξ /=.
-  apply N_β'. by asimpl. eauto.
+  apply sn_mutual => n; try qauto ctrs:SN, SNe, TRedSN depth:1.
+  move => */=. apply N_β';eauto. by asimpl.
 
-  move => * /=.
-  apply N_IndSuc';eauto. by asimpl.
+  move => */=. apply N_IndSuc'; eauto. by asimpl.
 Qed.
 
-Lemma ne_nf_embed n (a : PTm n) :
+Lemma ne_nf_embed (a : PTm) :
   (ne a -> SNe a) /\ (nf a -> SN a).
 Proof.
-  elim : n  / a => //=; hauto qb:on ctrs:SNe, SN.
+  elim :  a => //=; hauto qb:on ctrs:SNe, SN.
 Qed.
 
 #[export]Hint Constructors SN SNe TRedSN : sn.
@@ -383,49 +375,53 @@ Ltac2 rec solve_anti_ren () :=
   let x := Fresh.in_goal (Option.get (Ident.of_string "x")) in
   intro $x;
   lazy_match! Constr.type (Control.hyp x) with
-  | fin _ -> _ _ => (ltac1:(case;qauto depth:2 db:sn))
+  | nat -> nat => (ltac1:(case;qauto depth:2 db:sn))
+  | nat -> PTm => (ltac1:(case;qauto depth:2 db:sn))
   | _ => solve_anti_ren ()
   end.
 
 Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
 
-Lemma sn_antirenaming n :
-  (forall (a : PTm n) (s : SNe a), forall m (ξ : fin m -> fin n) b, a = ren_PTm ξ b -> SNe b) /\
-  (forall (a : PTm n) (s : SN a), forall m (ξ : fin m -> fin n) b, a = ren_PTm ξ b -> SN b) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall m (ξ : fin m -> fin n) a0,
+Lemma sn_antirenaming :
+  (forall (a : PTm) (s : SNe a), forall (ξ : nat -> nat) b, a = ren_PTm ξ b -> SNe b) /\
+  (forall (a : PTm) (s : SN a), forall (ξ : nat -> nat) b, a = ren_PTm ξ b -> SN b) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall (ξ : nat -> nat) a0,
       a = ren_PTm ξ a0 -> exists b0, TRedSN a0 b0 /\ b = ren_PTm ξ b0).
 Proof.
-  move : n. apply sn_mutual => n; try solve_anti_ren.
+  apply sn_mutual; try solve_anti_ren.
+  move => *. subst. spec_refl. hauto lq:on ctrs:TRedSN, SN.
 
-  move => a b ha iha m ξ []//= u u0 [+ ?]. subst.
+  move => a b ha iha ξ []//= u u0 [+ ?]. subst.
   case : u => //= => u [*]. subst.
   spec_refl. eexists. split. apply N_β=>//. by asimpl.
 
-  move => a b hb ihb m ξ[]//= p p0 [? +]. subst.
+  move => a b hb ihb ξ[]//= p p0 [? +]. subst.
   case : p0 => //= p p0 [*]. subst.
   spec_refl. by eauto with sn.
 
-  move => a b ha iha m ξ[]//= u u0 [? ]. subst.
+  move => a b ha iha ξ[]//= u u0 [? ]. subst.
   case : u0 => //=. move => p p0 [*]. subst.
   spec_refl. by eauto with sn.
 
-  move => P b c ha iha hb ihb hc ihc m ξ []//= P0 a0 b0 c0 [?]. subst.
+  move => P b c ha iha hb ihb hc ihc ξ []//= P0 a0 b0 c0 [?]. subst.
   case : a0 => //= _ *. subst.
   spec_refl. by eauto with sn.
 
-  move => P a b c hP ihP ha iha hb ihb hc ihc m ξ []//= P0 a0 b0 c0 [?]. subst.
+  move => P a b c hP ihP ha iha hb ihb hc ihc ξ []//= P0 a0 b0 c0 [?]. subst.
   case : a0 => //= a0 [*]. subst.
   spec_refl. eexists; repeat split; eauto with sn.
-  by asimpl. Qed.
+  by asimpl.
+Qed.
 
-Lemma sn_unmorphing n :
-  (forall (a : PTm n) (s : SNe a), forall m (ρ : fin m -> PTm n) b, a = subst_PTm ρ b -> SNe b) /\
-  (forall (a : PTm n) (s : SN a), forall m (ρ : fin m -> PTm n) b, a = subst_PTm ρ b -> SN b) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall m (ρ : fin m -> PTm n) a0,
+Lemma sn_unmorphing :
+  (forall (a : PTm) (s : SNe a), forall (ρ : nat -> PTm) b, a = subst_PTm ρ b -> SNe b) /\
+  (forall (a : PTm) (s : SN a), forall (ρ : nat -> PTm) b, a = subst_PTm ρ b -> SN b) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall (ρ : nat -> PTm) a0,
       a = subst_PTm ρ a0 -> (exists b0, b = subst_PTm ρ b0 /\ TRedSN a0 b0) \/ SNe a0).
 Proof.
-  move : n. apply sn_mutual => n; try solve_anti_ren.
-  - move => a b ha iha  m ξ b0.
+  apply sn_mutual; try solve_anti_ren.
+  - hauto q:on db:sn.
+  - move => a b ha iha  ξ b0.
     case : b0 => //=.
     + hauto lq:on rew:off db:sn.
     + move => p p0 [+ ?]. subst.
@@ -436,7 +432,7 @@ Proof.
       spec_refl.
       eexists. split; last by eauto using N_β.
       by asimpl.
-  - move => a0 a1 b hb ihb ha iha m ρ []//=.
+  - move => a0 a1 b hb ihb ha iha ρ []//=.
     + hauto lq:on rew:off db:sn.
     + move => t0  t1 [*]. subst.
       spec_refl.
@@ -447,18 +443,18 @@ Proof.
       * move => h.
         right.
         apply N_App => //.
-  - move => a b hb ihb m ρ []//=.
+  - move => a b hb ihb ρ []//=.
     + hauto l:on ctrs:TRedSN.
     + move => p p0 [?]. subst.
       case : p0 => //=.
       * hauto lq:on rew:off db:sn.
       * move => p p0 [*]. subst.
         hauto lq:on db:sn.
-  - move => a b ha iha m ρ []//=; first by hauto l:on db:sn.
+  - move => a b ha iha ρ []//=; first by hauto l:on db:sn.
     case => //=. move => []//=.
     + hauto lq:on db:sn.
     + hauto lq:on db:sn.
-  - move => p a b ha iha m ρ []//=; first by hauto l:on db:sn.
+  - move => p a b ha iha ρ []//=; first by hauto l:on db:sn.
     move => t0 t1 [*]. subst.
     spec_refl.
     case : iha.
@@ -466,12 +462,12 @@ Proof.
       left. eexists. split; last by eauto with sn.
       reflexivity.
     + hauto lq:on db:sn.
-  - move => P b c hP ihP hb ihb hc ihc m ρ []//=.
+  - move => P b c hP ihP hb ihb hc ihc ρ []//=.
     + hauto lq:on db:sn.
     + move => p []//=.
       * hauto lq:on db:sn.
       * hauto q:on db:sn.
-  - move => P a b c hP ihP ha iha hb ihb hc ihc m ρ []//=.
+  - move => P a b c hP ihP ha iha hb ihb hc ihc ρ []//=.
     + hauto lq:on db:sn.
     + move => P0 a0 b0 c0 [?]. subst.
       case : a0 => //=.
@@ -480,7 +476,7 @@ Proof.
         spec_refl.
         left. eexists. split; last by eauto with sn.
         by asimpl.
-  - move => P a0 a1 b c hP ihP hb ihb hc ihc ha iha m ρ[]//=.
+  - move => P a0 a1 b c hP ihP hb ihb hc ihc ha iha ρ[]//=.
     + hauto lq:on db:sn.
     + move => P1 a2 b2 c2 [*]. subst.
       spec_refl.
@@ -491,42 +487,41 @@ Proof.
       * hauto lq:on db:sn.
 Qed.
 
-Lemma SN_AppInv : forall n (a b : PTm n), SN (PApp a b) -> SN a /\ SN b.
+Lemma SN_AppInv : forall (a b : PTm), SN (PApp a b) -> SN a /\ SN b.
 Proof.
-  move => n a b. move E : (PApp a b) => u hu. move : a b E.
-  elim : n u /hu=>//=.
+  move => a b. move E : (PApp a b) => u hu. move : a b E.
+  elim : u /hu=>//=.
   hauto lq:on rew:off inv:SNe db:sn.
-  move => n a b ha hb ihb a0 b0 ?. subst.
+  move => a b ha hb ihb a0 b0 ?. subst.
   inversion ha; subst.
   move {ihb}.
   hecrush use:sn_unmorphing.
   hauto lq:on db:sn.
 Qed.
 
-Lemma SN_ProjInv : forall n p (a : PTm n), SN (PProj p a) -> SN a.
+Lemma SN_ProjInv : forall p (a : PTm), SN (PProj p a) -> SN a.
 Proof.
-  move => n p a. move E : (PProj p a) => u hu.
+  move => p a. move E : (PProj p a) => u hu.
   move : p a E.
-  elim : n u / hu => //=.
+  elim : u / hu => //=.
   hauto lq:on rew:off inv:SNe db:sn.
   hauto lq:on rew:off inv:TRedSN db:sn.
 Qed.
 
-Lemma SN_IndInv : forall n P (a : PTm n) b c, SN (PInd P a b c) -> SN P /\ SN a /\ SN b /\ SN c.
+Lemma SN_IndInv : forall P (a : PTm) b c, SN (PInd P a b c) -> SN P /\ SN a /\ SN b /\ SN c.
 Proof.
-  move => n P a b c. move E : (PInd P a b c) => u hu. move : P a b c E.
-  elim  : n u / hu => //=.
+  move => P a b c. move E : (PInd P a b c) => u hu. move : P a b c E.
+  elim  : u / hu => //=.
   hauto lq:on rew:off inv:SNe db:sn.
   hauto lq:on rew:off inv:TRedSN db:sn.
 Qed.
 
-Lemma epar_sn_preservation n :
-  (forall (a : PTm n) (s : SNe a), forall b, EPar.R a b -> SNe b) /\
-  (forall (a : PTm n) (s : SN a), forall b, EPar.R a b -> SN b) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall c, EPar.R a c -> exists d, TRedSN' c d /\ EPar.R b d).
+Lemma epar_sn_preservation :
+  (forall (a : PTm) (s : SNe a), forall b, EPar.R a b -> SNe b) /\
+  (forall (a : PTm) (s : SN a), forall b, EPar.R a b -> SN b) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall c, EPar.R a c -> exists d, TRedSN' c d /\ EPar.R b d).
 Proof.
-  move : n. apply sn_mutual => n.
-  - sauto lq:on.
+  apply sn_mutual.
   - sauto lq:on.
   - sauto lq:on.
   - sauto lq:on.
@@ -562,7 +557,7 @@ Proof.
       exists (subst_PTm (scons b1 VarPTm) a2).
       split.
       sauto lq:on.
-      hauto lq:on use:EPar.morphing, EPar.refl inv:option.
+      hauto lq:on use:EPar.morphing, EPar.refl inv:nat.
   - sauto.
   - move => a b hb ihb c.
     elim /EPar.inv => //= _.
@@ -590,12 +585,12 @@ Proof.
     elim /EPar.inv : ha0 => //=_.
     move => a0 a2 ha0 [*]. subst.
     eexists. split. apply T_Once. apply N_IndSuc; eauto.
-    hauto q:on ctrs:EPar.R use:EPar.morphing, EPar.refl inv:option.
+    hauto q:on ctrs:EPar.R use:EPar.morphing, EPar.refl inv:nat.
   - sauto q:on.
 Qed.
 
 Module RRed.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (****************** Beta ***********************)
   | AppAbs a b :
     R (PApp (PAbs a) b)  (subst_PTm (scons b VarPTm) a)
@@ -649,27 +644,27 @@ Module RRed.
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
 
-  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
+  Derive Inversion inv with (forall (a b : PTm), R a b) Sort Prop.
 
-  Lemma AppAbs' n a (b : PTm n) u :
+  Lemma AppAbs' a (b : PTm) u :
     u = (subst_PTm (scons b VarPTm) a) ->
     R (PApp (PAbs a) b)  u.
   Proof.
     move => ->. by apply AppAbs. Qed.
 
-  Lemma IndSuc' n P a b c u :
+  Lemma IndSuc' P a b c u :
     u = (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ->
-    R (@PInd n P (PSuc a) b c) u.
+    R (PInd P (PSuc a) b c) u.
   Proof. move => ->. apply IndSuc. 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.
 
     all : try qauto ctrs:R.
-    move => n a b m ξ /=.
+    move => a b ξ /=.
     apply AppAbs'. by asimpl.
     move => */=; apply IndSuc'; eauto. by asimpl.
   Qed.
@@ -678,56 +673,56 @@ Module RRed.
     let x := Fresh.in_goal (Option.get (Ident.of_string "x")) in
     intro $x;
     lazy_match! Constr.type (Control.hyp x) with
-    | fin _ -> _ _ => (ltac1:(case;hauto q:on depth:2 ctrs:RRed.R))
+    | nat -> nat => (ltac1:(case;hauto q:on depth:2 ctrs:RRed.R))
+    | nat -> PTm => (ltac1:(case;hauto q:on depth:2 ctrs:RRed.R))
     | _ => solve_anti_ren ()
     end.
 
   Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
 
-  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; try solve_anti_ren.
-    - move => n a b m ξ []//=. move => []//= t t0 [*]. subst.
+    move : ξ a E. elim : u b/h; try solve_anti_ren.
+    - move => a b ξ []//=. move => []//= t t0 [*]. subst.
       eexists. split. apply AppAbs. by asimpl.
-    - move => n p a b m ξ []//=.
+    - move => p a b ξ []//=.
       move => p0 []//=. hauto q:on ctrs:R.
-    - move => n p b c m ξ []//= P a0 b0 c0 [*]. subst.
+    - move => p b c ξ []//= P a0 b0 c0 [*]. subst.
       destruct a0 => //=.
       hauto lq:on ctrs:R.
-    - move => n P a b c m ξ []//=.
+    - move => P a b c ξ []//=.
       move => p p0 p1 p2 [?].  subst.
       case : p0 => //=.
       move => p0 [?] *. subst. eexists. split; eauto using IndSuc.
       by asimpl.
   Qed.
 
-  Lemma nf_imp n (a b : PTm n) :
+  Lemma nf_imp (a b : PTm) :
     nf a ->
     R a b -> False.
   Proof. move/[swap]. induction 1; hauto qb:on inv:PTm. Qed.
 
-  Lemma FromRedSN n (a b : PTm n) :
+  Lemma FromRedSN (a b : PTm) :
     TRedSN a b ->
     RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.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.
-    move => h. move : m ρ. elim : n a b / h => n.
+    move => h. move :  ρ. elim : a b / h => n.
 
     all : try hauto lq:on ctrs:R.
-    move => a b m ρ /=.
-    eapply AppAbs'; eauto; cycle 1. by asimpl.
+    move => */=. eapply AppAbs'; eauto; cycle 1. by asimpl.
     move => */=; apply : IndSuc'; eauto. by asimpl.
   Qed.
 
 End RRed.
 
 Module RPar.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (****************** Beta ***********************)
   | AppAbs a0 a1 b0 b1 :
     R a0 a1 ->
@@ -774,8 +769,6 @@ Module RPar.
     R A0 A1 ->
     R B0 B1 ->
     R (PBind p A0 B0) (PBind p A1 B1)
-  | BotCong :
-    R PBot PBot
   | NatCong :
     R PNat PNat
   | IndCong P0 P1 a0 a1 b0 b1 c0 c1 :
@@ -792,28 +785,28 @@ Module RPar.
     (* ------------ *)
     R (PSuc a0) (PSuc a1).
 
-  Lemma refl n (a : PTm n) : R a a.
+  Lemma refl (a : PTm) : R a a.
   Proof.
-    elim : n / a; hauto lq:on ctrs:R.
+    elim : a; hauto lq:on ctrs:R.
   Qed.
 
-  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
+  Derive Dependent Inversion inv with (forall (a b : PTm), R a b) Sort Prop.
 
-  Lemma AppAbs' n a0 a1 (b0 b1 : PTm n) u :
+  Lemma AppAbs' a0 a1 (b0 b1 : PTm) u :
     u = (subst_PTm (scons b1 VarPTm) a1) ->
     R a0 a1 ->
     R b0 b1 ->
     R (PApp (PAbs a0) b0)  u.
   Proof. move => ->. apply AppAbs. Qed.
 
-  Lemma ProjPair' n p u (a0 a1 b0 b1 : PTm n) :
+  Lemma ProjPair' p u (a0 a1 b0 b1 : PTm) :
     u = (if p is PL then a1 else b1) ->
     R a0 a1 ->
     R b0 b1 ->
     R (PProj p (PPair a0 b0)) u.
   Proof. move => ->. apply ProjPair. Qed.
 
-  Lemma IndSuc' n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 u :
+  Lemma IndSuc' P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 u :
     u = (subst_PTm (scons (PInd P1 a1 b1 c1) (scons a1 VarPTm)) c1) ->
     R P0 P1 ->
     R a0 a1 ->
@@ -823,43 +816,43 @@ Module RPar.
     R (PInd P0 (PSuc a0) b0 c0) u.
   Proof. move => ->. apply IndSuc. 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.
 
     all : try qauto ctrs:R use:ProjPair'.
 
-    move => n a0 a1 b0 b1 ha iha hb ihb m ξ /=.
+    move => a0 a1 b0 b1 ha iha hb ihb ξ /=.
     eapply AppAbs'; eauto. by asimpl.
 
     move => * /=. apply : IndSuc'; eauto. by asimpl.
   Qed.
 
-  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
+  Lemma morphing_ren (ρ0 ρ1 : nat -> PTm) (ξ : nat -> nat) :
     (forall i, R (ρ0 i) (ρ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 -> PTm m) a b  :
+  Lemma morphing_ext (ρ0 ρ1 : nat -> PTm) 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.
+  Proof. hauto q:on inv:nat. Qed.
 
-  Lemma morphing_up n m (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma morphing_up (ρ0 ρ1 : nat -> PTm) :
     (forall i, R (ρ0 i) (ρ1 i)) ->
     (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 : 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 => n.
+    move => + h. move : ρ0 ρ1. elim : a b / h.
     all : try hauto lq:on ctrs:R use:morphing_up, ProjPair'.
-    move => a0 a1 b0 b1 ha iha hb ihb m ρ0 ρ1 hρ /=.
+    move => a0 a1 b0 b1 ha iha hb ihb ρ0 ρ1 hρ /=.
     eapply AppAbs'; eauto; cycle 1. sfirstorder use:morphing_up. by asimpl.
     move => */=; eapply IndSuc'; eauto; cycle 1.
     sfirstorder use:morphing_up.
@@ -867,29 +860,29 @@ Module RPar.
     by asimpl.
   Qed.
 
-  Lemma substing n m (a : PTm n) b (ρ : fin n -> PTm m) :
+  Lemma substing (a : PTm) b (ρ : nat -> PTm) :
     R a b ->
     R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     hauto l:on use:morphing, refl.
   Qed.
 
-  Lemma cong n (a0 a1 : PTm (S n)) b0 b1  :
+  Lemma cong (a0 a1 : PTm) b0 b1  :
     R a0 a1 ->
     R b0 b1 ->
     R (subst_PTm (scons b0 VarPTm) a0) (subst_PTm (scons b1 VarPTm) a1).
   Proof.
     move => h0 h1. apply morphing=>//.
-    hauto q:on inv:option ctrs:R.
+    hauto q:on inv:nat ctrs:R.
   Qed.
 
-  Lemma FromRRed n (a b : PTm n) :
+  Lemma FromRRed (a b : PTm) :
     RRed.R a b -> RPar.R a b.
   Proof.
     induction 1; qauto l:on use:RPar.refl ctrs:RPar.R.
   Qed.
 
-  Function tstar {n} (a : PTm n) :=
+  Function tstar (a : PTm) :=
     match a with
     | VarPTm i => a
     | PAbs a => PAbs (tstar a)
@@ -900,7 +893,6 @@ Module RPar.
     | PProj p a => PProj p (tstar a)
     | PUniv i => PUniv i
     | PBind p A B => PBind p (tstar A) (tstar B)
-    | PBot => PBot
     | PNat => PNat
     | PZero => PZero
     | PSuc a => PSuc (tstar a)
@@ -910,11 +902,11 @@ Module RPar.
     | PInd P a b c => PInd (tstar P) (tstar a) (tstar b) (tstar c)
     end.
 
-  Lemma triangle n (a b : PTm n) :
+  Lemma triangle (a b : PTm) :
     RPar.R a b -> RPar.R b (tstar a).
   Proof.
     move : b.
-    apply tstar_ind => {}n{}a.
+    apply tstar_ind => {}a.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on rew:off inv:R use:cong ctrs:R.
@@ -942,33 +934,31 @@ Module RPar.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
-    - move => P b c ?. subst.
-      move => h0. inversion 1; subst. hauto lq:on ctrs:R. qauto l:on inv:R ctrs:R.
     - move => P a0 b c ? hP ihP ha iha hb ihb u. subst.
       elim /inv => //= _.
       + move => P0 P1 a1 a2 b0 b1 c0 c1 hP' ha' hb' hc' [*]. subst.
-        apply morphing. hauto lq:on ctrs:R inv:option.
+        apply morphing. hauto lq:on ctrs:R inv:nat.
         eauto.
       + sauto q:on ctrs:R.
     - sauto lq:on.
   Qed.
 
-  Lemma diamond n (a b c : PTm n) :
+  Lemma diamond (a b c : PTm) :
     R a b -> R a c -> exists d, R b d /\ R c d.
   Proof. eauto using triangle. Qed.
 End RPar.
 
-Lemma red_sn_preservation n :
-  (forall (a : PTm n) (s : SNe a), forall b, RPar.R a b -> SNe b) /\
-  (forall (a : PTm n) (s : SN a), forall b, RPar.R a b -> SN b) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall c, RPar.R a c -> exists d, TRedSN' c d /\ RPar.R b d).
+Lemma red_sn_preservation :
+  (forall (a : PTm) (s : SNe a), forall b, RPar.R a b -> SNe b) /\
+  (forall (a : PTm) (s : SN a), forall b, RPar.R a b -> SN b) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall c, RPar.R a c -> exists d, TRedSN' c d /\ RPar.R b d).
 Proof.
-  move : n. apply sn_mutual => n.
+  apply sn_mutual.
   - hauto l:on inv:RPar.R.
   - qauto l:on inv:RPar.R,SNe,SN ctrs:SNe.
   - hauto lq:on inv:RPar.R, SNe ctrs:SNe.
   - hauto lq:on inv:RPar.R, SNe ctrs:SNe.
-  - qauto l:on inv:RPar.R, SN,SNe ctrs:SNe.
+  (* - qauto l:on inv:RPar.R, SN,SNe ctrs:SNe. *)
   - qauto l:on ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN.
@@ -1002,13 +992,13 @@ Proof.
     elim /RPar.inv => //=_.
     + move => P0 P1 a0 a1 b0 b1 c0 c1 hP' ha' hb' hc' [*]. subst.
       eexists. split; first by apply T_Refl.
-      apply RPar.morphing=>//. hauto lq:on ctrs:RPar.R inv:option.
+      apply RPar.morphing=>//. hauto lq:on ctrs:RPar.R inv:nat.
     + move => P0 P1 a0 a1 b0 b1 c0 c1 hP' ha' hb' hc' [*]. subst.
       elim /RPar.inv : ha' => //=_.
       move => a0 a2 ha' [*]. subst.
       eexists. split. apply T_Once.
       apply N_IndSuc; eauto.
-      hauto q:on use:RPar.morphing ctrs:RPar.R inv:option.
+      hauto q:on use:RPar.morphing ctrs:RPar.R inv:nat.
   - sauto q:on.
 Qed.
 
@@ -1021,34 +1011,34 @@ Module RReds.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AbsCong n (a b : PTm (S n)) :
+  Lemma AbsCong (a b : PTm) :
     rtc RRed.R a b ->
     rtc RRed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma AppCong (a0 a1 b0 b1 : PTm) :
     rtc RRed.R a0 a1 ->
     rtc RRed.R b0 b1 ->
     rtc RRed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma PairCong (a0 a1 b0 b1 : PTm) :
     rtc RRed.R a0 a1 ->
     rtc RRed.R b0 b1 ->
     rtc RRed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm n) :
+  Lemma ProjCong p (a0 a1  : PTm) :
     rtc RRed.R a0 a1 ->
     rtc RRed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm n) :
+  Lemma SucCong (a0 a1 : PTm) :
     rtc RRed.R a0 a1 ->
     rtc RRed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 :
+  Lemma IndCong P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc RRed.R P0 P1 ->
     rtc RRed.R a0 a1 ->
     rtc RRed.R b0 b1 ->
@@ -1056,27 +1046,27 @@ Module RReds.
     rtc RRed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+  Lemma BindCong p (A0 A1 : PTm) B0 B1 :
     rtc RRed.R A0 A1 ->
     rtc RRed.R B0 B1 ->
     rtc RRed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
-  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+  Lemma renaming (a b : PTm) (ξ : nat -> nat) :
     rtc RRed.R a b -> rtc RRed.R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
-    move => h. move : m ξ. elim : a b /h; hauto lq:on ctrs:rtc use:RRed.renaming.
+    move => h. move : ξ. elim : a b /h; hauto lq:on ctrs:rtc use:RRed.renaming.
   Qed.
 
-  Lemma FromRPar n (a b : PTm n) (h : RPar.R a b) :
+  Lemma FromRPar (a b : PTm) (h : RPar.R a b) :
     rtc RRed.R a b.
   Proof.
-    elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong, IndCong, SucCong.
-    move => n a0 a1 b0 b1 ha iha hb ihb.
+    elim : a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong, IndCong, SucCong.
+    move => a0 a1 b0 b1 ha iha hb ihb.
     apply : rtc_r; last by apply RRed.AppAbs.
     by eauto using AppCong, AbsCong.
 
-    move => n p a0 a1 b0 b1 ha iha hb ihb.
+    move => p a0 a1 b0 b1 ha iha hb ihb.
     apply : rtc_r; last by apply RRed.ProjPair.
     by eauto using PairCong, ProjCong.
 
@@ -1087,7 +1077,7 @@ Module RReds.
     by eauto using SucCong, IndCong.
   Qed.
 
-  Lemma RParIff n (a b : PTm n) :
+  Lemma RParIff (a b : PTm) :
     rtc RRed.R a b <-> rtc RPar.R a b.
   Proof.
     split.
@@ -1095,11 +1085,11 @@ Module RReds.
     induction 1; hauto l:on ctrs:rtc use:FromRPar, @relations.rtc_transitive.
   Qed.
 
-  Lemma nf_refl n (a b : PTm n) :
+  Lemma nf_refl (a b : PTm) :
     rtc RRed.R a b -> nf a -> a = b.
   Proof. induction 1; sfirstorder use:RRed.nf_imp. Qed.
 
-  Lemma FromRedSNs n (a b : PTm n) :
+  Lemma FromRedSNs (a b : PTm) :
     rtc TRedSN a b ->
     rtc RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:RRed.FromRedSN. Qed.
@@ -1108,7 +1098,7 @@ End RReds.
 
 
 Module NeEPar.
-  Inductive R_nonelim {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R_nonelim : PTm -> PTm ->  Prop :=
   (****************** Eta ***********************)
   | AppEta a0 a1 :
     ~~ ishf a0 ->
@@ -1141,8 +1131,6 @@ Module NeEPar.
     R_nonelim A0 A1 ->
     R_nonelim B0 B1 ->
     R_nonelim (PBind p A0 B0) (PBind p A1 B1)
-  | BotCong :
-    R_nonelim PBot PBot
   | NatCong :
     R_nonelim PNat PNat
   | IndCong P0 P1 a0 a1 b0 b1 c0 c1 :
@@ -1158,7 +1146,7 @@ Module NeEPar.
     R_nonelim a0 a1 ->
     (* ------------ *)
     R_nonelim (PSuc a0) (PSuc a1)
-  with R_elim {n} : PTm n -> PTm n -> Prop :=
+  with R_elim : PTm -> PTm -> Prop :=
   | NAbsCong a0 a1 :
     R_nonelim a0 a1 ->
     R_elim (PAbs a0) (PAbs a1)
@@ -1181,8 +1169,6 @@ Module NeEPar.
     R_nonelim A0 A1 ->
     R_nonelim B0 B1 ->
     R_elim (PBind p A0 B0) (PBind p A1 B1)
-  | NBotCong :
-    R_elim PBot PBot
   | NNatCong :
     R_elim PNat PNat
   | NIndCong P0 P1 a0 a1 b0 b1 c0 c1 :
@@ -1204,11 +1190,11 @@ Module NeEPar.
 
   Combined Scheme epar_mutual from epar_elim_ind, epar_nonelim_ind.
 
-  Lemma R_elim_nf n :
-    (forall (a b : PTm n), R_elim a b -> nf b -> nf a) /\
-      (forall (a b : PTm n), R_nonelim a b -> nf b -> nf a).
+  Lemma R_elim_nf :
+    (forall (a b : PTm), R_elim a b -> nf b -> nf a) /\
+      (forall (a b : PTm), R_nonelim a b -> nf b -> nf a).
   Proof.
-    move : n. apply epar_mutual => n //=.
+    apply epar_mutual => //=.
     - move => a0 a1 b0 b1 h ih h' ih' /andP [h0 h1].
       have hb0 : nf b0 by eauto.
       suff : ne a0 by qauto b:on.
@@ -1237,22 +1223,22 @@ Module NeEPar.
     - hauto lqb:on inv:R_elim.
   Qed.
 
-  Lemma R_nonelim_nothf n (a b : PTm n) :
+  Lemma R_nonelim_nothf (a b : PTm) :
     R_nonelim a b ->
     ~~ ishf a ->
     R_elim a b.
   Proof.
-    move => h. elim : n a b /h=>//=; hauto lq:on ctrs:R_elim.
+    move => h. elim : a b /h=>//=; hauto lq:on ctrs:R_elim.
   Qed.
 
-  Lemma R_elim_nonelim n (a b : PTm n) :
+  Lemma R_elim_nonelim (a b : PTm) :
     R_elim a b ->
     R_nonelim a b.
-    move => h. elim : n a b /h=>//=; hauto lq:on ctrs:R_nonelim.
+    move => h. elim : a b /h=>//=; hauto lq:on ctrs:R_nonelim.
   Qed.
 
-  Lemma ToEPar : forall n, (forall (a b : PTm n), R_elim a b -> EPar.R a b) /\
-                 (forall (a b : PTm n), R_nonelim a b -> EPar.R a b).
+  Lemma ToEPar : (forall (a b : PTm), R_elim a b -> EPar.R a b) /\
+                 (forall (a b : PTm), R_nonelim a b -> EPar.R a b).
   Proof.
     apply epar_mutual; qauto l:on ctrs:EPar.R.
   Qed.
@@ -1260,51 +1246,45 @@ Module NeEPar.
 End NeEPar.
 
 Module Type NoForbid.
-  Parameter P : forall n, PTm n -> Prop.
-  Arguments P {n}.
+  Parameter P : PTm -> Prop.
 
-  Axiom P_EPar : forall n (a b : PTm n), EPar.R a b -> P a -> P b.
-  Axiom P_RRed : forall n (a b : PTm n), RRed.R a b -> P a -> P b.
-  (* Axiom P_AppPair : forall n (a b c : PTm n), ~ P (PApp (PPair a b) c). *)
-  (* Axiom P_ProjAbs : forall n p (a : PTm (S n)), ~ P (PProj p (PAbs a)). *)
-  (* Axiom P_ProjBind : forall n p p' (A : PTm n) B, ~ P (PProj p (PBind p' A B)). *)
-  (* Axiom P_AppBind : forall n p (A : PTm n) B b, ~ P (PApp (PBind p A B) b). *)
-  Axiom PApp_imp : forall n a b, @ishf n a -> ~~ isabs a -> ~ P (PApp a b).
-  Axiom PProj_imp : forall  n p a, @ishf n a -> ~~ ispair a -> ~ P (PProj p a).
-  Axiom PInd_imp : forall n Q (a : PTm n) b c,
+  Axiom P_EPar : forall (a b : PTm), EPar.R a b -> P a -> P b.
+  Axiom P_RRed : forall (a b : PTm), RRed.R a b -> P a -> P b.
+  Axiom PApp_imp : forall a b, ishf a -> ~~ isabs a -> ~ P (PApp a b).
+  Axiom PProj_imp : forall p a, ishf a -> ~~ ispair a -> ~ P (PProj p a).
+  Axiom PInd_imp : forall Q (a : PTm) b c,
   ishf a ->
   ~~ iszero a ->
   ~~ issuc a  -> ~ P (PInd Q a b c).
 
-  Axiom P_AppInv : forall n (a b : PTm n), P (PApp a b) -> P a /\ P b.
-  Axiom P_AbsInv : forall n (a : PTm (S n)), P (PAbs a) -> P a.
-  Axiom P_SucInv : forall n (a : PTm n), P (PSuc a) -> P a.
-  Axiom P_BindInv : forall n p (A : PTm n) B, P (PBind p A B) -> P A /\ P B.
-  Axiom P_IndInv : forall n Q (a : PTm n) b c, P (PInd Q a b c) -> P Q /\ P a /\ P b /\ P c.
-
-  Axiom P_PairInv : forall n (a b : PTm n), P (PPair a b) -> P a /\ P b.
-  Axiom P_ProjInv : forall n p (a : PTm n), P (PProj p a) -> P a.
-  Axiom P_renaming : forall n m (ξ : fin n -> fin m) a , P (ren_PTm ξ a) <-> P a.
+  Axiom P_AppInv : forall (a b : PTm), P (PApp a b) -> P a /\ P b.
+  Axiom P_AbsInv : forall (a : PTm), P (PAbs a) -> P a.
+  Axiom P_SucInv : forall (a : PTm), P (PSuc a) -> P a.
+  Axiom P_BindInv : forall p (A : PTm) B, P (PBind p A B) -> P A /\ P B.
+  Axiom P_IndInv : forall Q (a : PTm) b c, P (PInd Q a b c) -> P Q /\ P a /\ P b /\ P c.
 
+  Axiom P_PairInv : forall (a b : PTm), P (PPair a b) -> P a /\ P b.
+  Axiom P_ProjInv : forall p (a : PTm), P (PProj p a) -> P a.
+  Axiom P_renaming : forall (ξ : nat -> nat) a , P (ren_PTm ξ a) <-> P a.
 End NoForbid.
 
 Module Type NoForbid_FactSig (M : NoForbid).
 
-  Axiom P_EPars : forall n (a b : PTm n), rtc EPar.R a b -> M.P a -> M.P b.
+  Axiom P_EPars : forall (a b : PTm), rtc EPar.R a b -> M.P a -> M.P b.
 
-  Axiom P_RReds : forall n (a b : PTm n), rtc RRed.R a b -> M.P a -> M.P b.
+  Axiom P_RReds : forall (a b : PTm), rtc RRed.R a b -> M.P a -> M.P b.
 
 End NoForbid_FactSig.
 
 Module NoForbid_Fact (M : NoForbid) : NoForbid_FactSig M.
   Import M.
 
-  Lemma P_EPars : forall n (a b : PTm n), rtc EPar.R a b -> P a -> P b.
+  Lemma P_EPars : forall (a b : PTm), rtc EPar.R a b -> P a -> P b.
   Proof.
     induction 1; eauto using P_EPar, rtc_l, rtc_refl.
   Qed.
 
-  Lemma P_RReds : forall n (a b : PTm n), rtc RRed.R a b -> P a -> P b.
+  Lemma P_RReds : forall (a b : PTm), rtc RRed.R a b -> P a -> P b.
   Proof.
     induction 1; eauto using P_RRed, rtc_l, rtc_refl.
   Qed.
@@ -1312,62 +1292,61 @@ End NoForbid_Fact.
 
 Module SN_NoForbid <: NoForbid.
   Definition P := @SN.
-  Arguments P {n}.
 
-  Lemma P_EPar : forall n (a b : PTm n), EPar.R a b -> P a -> P b.
+  Lemma P_EPar : forall (a b : PTm), EPar.R a b -> P a -> P b.
   Proof. sfirstorder use:epar_sn_preservation. Qed.
 
-  Lemma P_RRed : forall n (a b : PTm n), RRed.R a b -> P a -> P b.
+  Lemma P_RRed : forall (a b : PTm), RRed.R a b -> P a -> P b.
   Proof. hauto q:on use:red_sn_preservation, RPar.FromRRed. Qed.
 
-  Lemma PApp_imp : forall n a b, @ishf n a -> ~~ isabs a -> ~ P (PApp a b).
+  Lemma PApp_imp : forall a b, ishf a -> ~~ isabs a -> ~ P (PApp a b).
     sfirstorder use:fp_red.PApp_imp. Qed.
-  Lemma PProj_imp : forall  n p a, @ishf n a -> ~~ ispair a -> ~ P (PProj p a).
+  Lemma PProj_imp : forall p a, ishf a -> ~~ ispair a -> ~ P (PProj p a).
     sfirstorder use:fp_red.PProj_imp. Qed.
 
-  Lemma PInd_imp : forall n Q (a : PTm n) b c,
+  Lemma PInd_imp : forall Q (a : PTm) b c,
       ishf a ->
       ~~ iszero a ->
       ~~ issuc a  -> ~ P (PInd Q a b c).
   Proof. sfirstorder use: fp_red.PInd_imp. Qed.
 
-  Lemma P_AppInv : forall n (a b : PTm n), P (PApp a b) -> P a /\ P b.
+  Lemma P_AppInv : forall (a b : PTm), P (PApp a b) -> P a /\ P b.
   Proof. sfirstorder use:SN_AppInv. Qed.
 
-  Lemma P_PairInv : forall n (a b : PTm n), P (PPair a b) -> P a /\ P b.
-    move => n a b. move E : (PPair a b) => u h.
-    move : a b E. elim : n u / h; sauto lq:on rew:off. Qed.
+  Lemma P_PairInv : forall (a b : PTm), P (PPair a b) -> P a /\ P b.
+    move => a b. move E : (PPair a b) => u h.
+    move : a b E. elim : u / h; sauto lq:on rew:off. Qed.
 
-  Lemma P_ProjInv : forall n p (a : PTm n), P (PProj p a) -> P a.
+  Lemma P_ProjInv : forall p (a : PTm), P (PProj p a) -> P a.
   Proof. sfirstorder use:SN_ProjInv. Qed.
 
-  Lemma P_BindInv : forall n p (A : PTm n) B, P (PBind p A B) -> P A /\ P B.
+  Lemma P_BindInv : forall p (A : PTm) B, P (PBind p A B) -> P A /\ P B.
   Proof.
-    move => n p A B.
+    move => p A B.
     move E : (PBind p A B) => u hu.
-    move : p A B E. elim : n u /hu=>//=;sauto lq:on rew:off.
+    move : p A B E. elim : u /hu=>//=;sauto lq:on rew:off.
   Qed.
 
-  Lemma P_SucInv : forall n (a : PTm n), P (PSuc a) -> P a.
+  Lemma P_SucInv : forall (a : PTm), P (PSuc a) -> P a.
   Proof. sauto lq:on. Qed.
 
-  Lemma P_AbsInv : forall n (a : PTm (S n)), P (PAbs a) -> P a.
+  Lemma P_AbsInv : forall (a : PTm), P (PAbs a) -> P a.
   Proof.
-    move => n a. move E : (PAbs a) => u h.
+    move => a. move E : (PAbs a) => u h.
     move : E. move : a.
     induction h; sauto lq:on rew:off.
   Qed.
 
-  Lemma P_renaming : forall n m (ξ : fin n -> fin m) a , P (ren_PTm ξ a) <-> P  a.
+  Lemma P_renaming : forall (ξ : nat -> nat) a , P (ren_PTm ξ a) <-> P  a.
   Proof. hauto lq:on use:sn_antirenaming, sn_renaming. Qed.
 
-  Lemma P_ProjBind : forall n p p' (A : PTm n) B, ~ P (PProj p (PBind p' A B)).
+  Lemma P_ProjBind : forall p p' (A : PTm) B, ~ P (PProj p (PBind p' A B)).
   Proof. sfirstorder use:PProjBind_imp. Qed.
 
-  Lemma P_AppBind : forall n p (A : PTm n) B b, ~ P (PApp (PBind p A B) b).
+  Lemma P_AppBind : forall p (A : PTm) B b, ~ P (PApp (PBind p A B) b).
   Proof. sfirstorder use:PAppBind_imp. Qed.
 
-  Lemma P_IndInv : forall n Q (a : PTm n) b c, P (PInd Q a b c) -> P Q /\ P a /\ P b /\ P c.
+  Lemma P_IndInv : forall Q (a : PTm) b c, P (PInd Q a b c) -> P Q /\ P a /\ P b /\ P c.
   Proof. sfirstorder use:SN_IndInv. Qed.
 
 End SN_NoForbid.
@@ -1378,13 +1357,13 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
   Import M MFacts.
   #[local]Hint Resolve P_EPar P_RRed PApp_imp PProj_imp : forbid.
 
-  Lemma η_split n (a0 a1 : PTm n) :
+  Lemma η_split (a0 a1 : PTm) :
     EPar.R a0 a1 ->
     P a0 ->
     exists b, rtc RRed.R a0 b /\ NeEPar.R_nonelim b a1.
   Proof.
-    move => h. elim : n a0 a1 /h .
-    - move => n a0 a1 ha ih /[dup] hP.
+    move => h. elim : a0 a1 /h .
+    - move => a0 a1 ha ih /[dup] hP.
       move /P_AbsInv /P_AppInv => [/P_renaming ha0 _].
       have {ih} := ih ha0.
       move => [b [ih0 ih1]].
@@ -1404,7 +1383,13 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         by eauto using RReds.renaming.
         apply rtc_refl.
         apply : RRed.AbsCong => /=.
-        apply RRed.AppAbs'. by asimpl.
+        apply RRed.AppAbs'. asimpl.
+        set y := subst_PTm _ _.
+        suff : ren_PTm id p = y. by asimpl.
+        subst y.
+        substify.
+        apply ext_PTm.
+        case => //=.
       (* violates SN *)
       + move /P_AbsInv in hP.
         have {}hP : P (PApp (ren_PTm shift b) (VarPTm var_zero))
@@ -1414,7 +1399,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         have ? : ishf (ren_PTm shift b) by scongruence use:ishf_ren.
         exfalso.
         sfirstorder use:PApp_imp.
-    - move => n a0 a1 h ih /[dup] hP.
+    - move => a0 a1 h ih /[dup] hP.
       move /P_PairInv => [/P_ProjInv + _].
       move : ih => /[apply].
       move => [b [ih0 ih1]].
@@ -1439,7 +1424,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         move : P_RReds hP. repeat move/[apply] => /=.
         sfirstorder use:PProj_imp.
     - hauto lq:on ctrs:NeEPar.R_nonelim use:RReds.AbsCong, P_AbsInv.
-    - move => n a0 a1 b0 b1 ha iha hb ihb.
+    - move => a0 a1 b0 b1 ha iha hb ihb.
       move => /[dup] hP /P_AppInv [hP0 hP1].
       have {iha} [a2 [iha0 iha1]]  := iha hP0.
       have {ihb} [b2 [ihb0 ihb1]]  := ihb hP1.
@@ -1462,7 +1447,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           have : P (PApp a2 b0) by sfirstorder use:RReds.AppCong, @rtc_refl, P_RReds.
           sfirstorder use:PApp_imp.
     - hauto lq:on ctrs:NeEPar.R_nonelim use:RReds.PairCong, P_PairInv.
-    - move => n p a0 a1 ha ih /[dup] hP /P_ProjInv.
+    - move => p a0 a1 ha ih /[dup] hP /P_ProjInv.
       move : ih => /[apply]. move => [a2 [iha0 iha1]].
       case /orP : (orNb (ishf a2)) => [h|].
       exists (PProj p a2).
@@ -1487,8 +1472,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto l:on.
     - hauto lq:on ctrs:NeEPar.R_nonelim, rtc use:RReds.BindCong, P_BindInv.
     - hauto lq:on ctrs:NeEPar.R_nonelim, rtc use:RReds.BindCong, P_BindInv.
-    - hauto l:on ctrs:NeEPar.R_nonelim.
-    - move => n P0 P1 a0 a1 b0 b1 c0 c1 hP ihP ha iha hb ihb hc ihc /[dup] hInd /P_IndInv.
+    - move => P0 P1 a0 a1 b0 b1 c0 c1 hP ihP ha iha hb ihb hc ihc /[dup] hInd /P_IndInv.
       move => []. move : ihP => /[apply].
       move => [P01][ihP0]ihP1.
       move => []. move : iha => /[apply].
@@ -1514,32 +1498,32 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto lq:on ctrs:NeEPar.R_nonelim use:RReds.SucCong, P_SucInv.
   Qed.
 
-  Lemma eta_postponement n a b c :
-    @P n a ->
+  Lemma eta_postponement a b c :
+    P a ->
     EPar.R a b ->
     RRed.R b c ->
     exists d, rtc RRed.R a d /\ EPar.R d c.
   Proof.
     move => + h.
     move : c.
-    elim : n a b /h => //=.
-    - move => n a0 a1 ha iha c /[dup] hP /P_AbsInv /P_AppInv [/P_renaming hP' _] hc.
+    elim : a b /h => //=.
+    - move => a0 a1 ha iha c /[dup] hP /P_AbsInv /P_AppInv [/P_renaming hP' _] hc.
       move : iha (hP') (hc); repeat move/[apply].
       move => [d [h0 h1]].
       exists  (PAbs (PApp (ren_PTm shift d) (VarPTm var_zero))).
       split. hauto lq:on rew:off ctrs:rtc use:RReds.AbsCong, RReds.AppCong, RReds.renaming.
       hauto lq:on ctrs:EPar.R.
-    - move => n a0 a1 ha iha c /P_PairInv [/P_ProjInv + _].
+    - move => a0 a1 ha iha c /P_PairInv [/P_ProjInv + _].
       move /iha => /[apply].
       move => [d [h0 h1]].
       exists (PPair (PProj PL d) (PProj PR d)).
       hauto lq:on ctrs:EPar.R use:RReds.PairCong, RReds.ProjCong.
-    - move => n a0 a1 ha iha c  /P_AbsInv /[swap].
+    - move => a0 a1 ha iha c  /P_AbsInv /[swap].
       elim /RRed.inv => //=_.
       move => a2 a3 + [? ?]. subst.
       move : iha; repeat move/[apply].
       hauto lq:on use:RReds.AbsCong ctrs:EPar.R.
-    - move => n a0 a1 b0 b1 ha iha hb ihb c hP.
+    - move => a0 a1 b0 b1 ha iha hb ihb c hP.
       elim /RRed.inv => //= _.
       + move => a2 b2 [*]. subst.
         have [hP' hP''] : P a0 /\ P b0 by sfirstorder use:P_AppInv.
@@ -1558,7 +1542,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           apply RRed.AppCong0. apply RRed.AbsCong. simpl. apply RRed.AppAbs.
           asimpl.
           apply rtc_once.
-          apply RRed.AppAbs.
+          apply RRed.AppAbs'. by asimpl.
         * exfalso.
           move : hP h0. clear => hP h0.
           have : rtc RRed.R (PApp a0 b0) (PApp (PPair (PProj PL a1) (PProj PR a1)) b0)
@@ -1569,7 +1553,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           split.
           apply : rtc_r; last by apply RRed.AppAbs.
           hauto lq:on ctrs:rtc use:RReds.AppCong.
-          hauto l:on inv:option use:EPar.morphing,NeEPar.ToEPar.
+          hauto l:on inv:nat use:EPar.morphing,NeEPar.ToEPar.
       + move => a2 a3 b2 ha2 [*]. subst.
         move : iha (ha2) {ihb} => /[apply].
         have : P a0 by sfirstorder use:P_AppInv.
@@ -1582,10 +1566,10 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         have : P b0 by sfirstorder use:P_AppInv.
         move : ihb hb2; repeat move /[apply].
         hauto lq:on rew:off ctrs:EPar.R, rtc use:RReds.AppCong.
-    - move => n a0 a1 b0 b1 ha iha hb ihb c /P_PairInv [hP hP'].
+    - move => a0 a1 b0 b1 ha iha hb ihb c /P_PairInv [hP hP'].
       elim /RRed.inv => //=_;
         hauto lq:on rew:off ctrs:EPar.R, rtc use:RReds.PairCong.
-    - move => n p a0 a1 ha iha c /[dup] hP /P_ProjInv hP'.
+    - move => p a0 a1 ha iha c /[dup] hP /P_ProjInv hP'.
       elim / RRed.inv => //= _.
       + move => p0 a2 b0 [*]. subst.
         move : η_split hP'  ha; repeat move/[apply].
@@ -1619,8 +1603,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto lq:on inv:RRed.R ctrs:rtc.
     - sauto lq:on ctrs:EPar.R, rtc use:RReds.BindCong, P_BindInv, @relations.rtc_transitive.
     - hauto lq:on inv:RRed.R ctrs:rtc.
-    - hauto q:on ctrs:rtc inv:RRed.R.
-    - move => n P0 P1 a0 a1 b0 b1 c0 c1 hP ihP ha iha hb ihb hc ihc u.
+    - move => P0 P1 a0 a1 b0 b1 c0 c1 hP ihP ha iha hb ihb hc ihc u.
       move => /[dup] hInd.
       move /P_IndInv.
       move => [pP0][pa0][pb0]pc0.
@@ -1650,7 +1633,12 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           apply : rtc_r.
           apply RReds.IndCong; eauto; eauto using rtc_refl.
           apply RRed.IndSuc.
-          apply EPar.morphing;eauto. hauto lq:on ctrs:EPar.R inv:option use:NeEPar.ToEPar.
+          apply EPar.morphing;eauto.
+          case => [|].
+          hauto lq:on rew:off ctrs:EPar.R use:NeEPar.ToEPar.
+          case => [|i].
+          hauto lq:on rew:off ctrs:EPar.R use:NeEPar.ToEPar.
+          asimpl. apply EPar.VarTm.
       + move => P2 P3 a2 b2 c hP0 [*]. subst.
         move : ihP hP0 pP0. repeat move/[apply].
         move => [P2][h0]h1.
@@ -1672,7 +1660,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         exists (PInd P0 a0 b0 c2).
         sfirstorder use:RReds.IndCong, @rtc_refl, EPar.IndCong.
     - hauto lq:on inv:RRed.R ctrs:rtc, EPar.R.
-    - move => n a0 a1 ha iha u /P_SucInv ha0.
+    - move => a0 a1 ha iha u /P_SucInv ha0.
       elim /RRed.inv => //= _ a2 a3 h [*]. subst.
       move : iha (ha0) (h); repeat move/[apply].
       move => [a2 [ih0 ih1]].
@@ -1681,8 +1669,8 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       by apply EPar.SucCong.
   Qed.
 
-  Lemma η_postponement_star n a b c :
-    @P n a ->
+  Lemma η_postponement_star a b c :
+    P a ->
     EPar.R a b ->
     rtc RRed.R b c ->
     exists d, rtc RRed.R a d /\ EPar.R d c.
@@ -1698,8 +1686,8 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       sfirstorder use:@relations.rtc_transitive.
   Qed.
 
-  Lemma η_postponement_star' n a b c :
-    @P n a ->
+  Lemma η_postponement_star' a b c :
+    P a ->
     EPar.R a b ->
     rtc RRed.R b c ->
     exists d, rtc RRed.R a d /\ NeEPar.R_nonelim d c.
@@ -1716,7 +1704,7 @@ End UniqueNF.
 Module SN_UniqueNF := UniqueNF SN_NoForbid NoForbid_FactSN.
 
 Module ERed.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
 
   (****************** Eta ***********************)
   | AppEta a :
@@ -1765,9 +1753,9 @@ Module ERed.
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
 
-  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
+  Derive Inversion inv with (forall (a b : PTm), R a b) Sort Prop.
 
-  Lemma ToEPar n (a b : PTm n) :
+  Lemma ToEPar (a b : PTm) :
     ERed.R a b -> EPar.R a b.
   Proof.
     induction 1; hauto lq:on use:EPar.refl ctrs:EPar.R.
@@ -1777,38 +1765,34 @@ Module ERed.
     let x := Fresh.in_goal (Option.get (Ident.of_string "x")) in
     intro $x;
     lazy_match! Constr.type (Control.hyp x) with
-    | fin _ -> _ _ => (ltac1:(case;hauto q:on depth:2 ctrs:ERed.R))
+    | nat -> _ => (ltac1:(case;hauto q:on depth:2 ctrs:ERed.R))
     | _ => solve_anti_ren ()
     end.
 
   Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
 
-  (* Definition down n m (ξ : fin n -> fin m) (a : fin (S n)) : fin m. *)
-  (*   destruct a. *)
-  (*   exact (ξ f). *)
-
-  Lemma AppEta' n a u :
-    u = (@PApp (S n) (ren_PTm shift a) (VarPTm var_zero)) ->
+  Lemma AppEta' a u :
+    u = (PApp (ren_PTm shift a) (VarPTm var_zero)) ->
     R (PAbs u) a.
   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 => n a m ξ /=.
+    move => a ξ /=.
     apply AppEta'; eauto. by asimpl.
     all : qauto ctrs:R.
   Qed.
 
   (* Characterization of variable freeness conditions *)
-  Definition tm_i_free {n} a (i : fin n) := exists m (ξ ξ0 : fin n -> fin m), ξ i <> ξ0 i /\ ren_PTm ξ a = ren_PTm ξ0 a.
+  Definition tm_i_free a (i : fin n) := exists m (ξ ξ0 : fin n -> fin m), ξ i <> ξ0 i /\ ren_PTm a = ren_PTm0 a.
 
   Lemma subst_differ_one_ren_up n m i (ξ0 ξ1 : fin n -> fin m) :
     (forall j, i <> j -> ξ0 j = ξ1 j) ->
-    (forall j, shift i <> j -> upRen_PTm_PTm ξ0 j =  upRen_PTm_PTm ξ1 j).
+    (forall j, shift i <> j -> upRen_PTm_PTm0 j =  upRen_PTm_PTm1 j).
   Proof.
     move => hξ.
     destruct j. asimpl.
@@ -1820,7 +1804,7 @@ Module ERed.
   Lemma tm_free_ren_any n a i :
     tm_i_free a i ->
     forall m (ξ0 ξ1 : fin n -> fin m), (forall j, i <> j -> ξ0 j = ξ1 j) ->
-             ren_PTm ξ0 a = ren_PTm ξ1 a.
+             ren_PTm0 a = ren_PTm1 a.
   Proof.
     rewrite /tm_i_free.
     move => [+ [+ [+ +]]].
@@ -1841,7 +1825,7 @@ Module ERed.
       move /subst_differ_one_ren_up in hξ.
       move /(_ (shift i)) in iha.
       move : iha hξ. repeat move/[apply].
-      move /(_ _ (upRen_PTm_PTm ρ0) (upRen_PTm_PTm ρ1)).
+      move /(_ _ (upRen_PTm_PTm0) (upRen_PTm_PTm1)).
       apply. hauto l:on.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
@@ -1851,12 +1835,12 @@ Module ERed.
     - admit.
   Admitted.
 
-  Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
+  Lemma antirenaming n m (a : PTm) (b : PTm) (ξ : fin n -> fin m) :
     (forall i j, ξ i = ξ j -> i = j) ->
-    R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
+    R (ren_PTm a) b -> exists b0, R a b0 /\ ren_PTm b0 = b.
   Proof.
     move => hξ.
-    move E : (ren_PTm ξ a) => u hu.
+    move E : (ren_PTm a) => u hu.
     move : n ξ a hξ E.
     elim : m u b / hu; try solve_anti_ren.
     - move => n a m ξ []//=.
@@ -1870,7 +1854,7 @@ Module ERed.
       move => [p ?]. subst.
       move : h. asimpl.
       replace (ren_PTm (funcomp shift ξ) p) with
-        (ren_PTm shift (ren_PTm ξ p)); last by asimpl.
+        (ren_PTm shift (ren_PTm p)); last by asimpl.
       move /ren_injective.
       move /(_ ltac:(hauto l:on)).
       move => ?. subst.
@@ -1886,14 +1870,14 @@ Module ERed.
       sauto lq:on use:up_injective.
     - move => n p A B0 B1 hB ihB m ξ + hξ.
       case => //= p' A2 B2 [*]. subst.
-      have : (forall i j, (upRen_PTm_PTm ξ) i = (upRen_PTm_PTm ξ) j -> i = j) by sauto.
+      have : (forall i j, (upRen_PTm_PTm) i = (upRen_PTm_PTm) j -> i = j) by sauto.
       move => {}/ihB => ihB.
       spec_refl.
       sauto lq:on.
   Admitted.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    R a b -> R (subst_PTm a) (subst_PTm b).
   Proof.
     move => h. move : m ρ. elim : n a b /h => n.
     move => a m ρ /=.
@@ -1901,7 +1885,7 @@ Module ERed.
     all : hauto lq:on ctrs:R.
   Qed.
 
-  Lemma nf_preservation n (a b : PTm n) :
+  Lemma nf_preservation n (a b : PTm) :
     ERed.R a b -> (nf a -> nf b) /\ (ne a -> ne b).
   Proof.
     move => h.
@@ -1925,36 +1909,36 @@ Module EReds.
     rtc ERed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
     rtc ERed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
     rtc ERed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm n) :
+  Lemma ProjCong n p (a0 a1  : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
     rtc ERed.R A0 A1 ->
     rtc ERed.R B0 B1 ->
     rtc ERed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
 
-  Lemma SucCong n (a0 a1 : PTm n) :
+  Lemma SucCong n (a0 a1 : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 :
+  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc ERed.R P0 P1 ->
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
@@ -1962,11 +1946,11 @@ Module EReds.
     rtc ERed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
-    rtc ERed.R a b -> rtc ERed.R (ren_PTm ξ a) (ren_PTm ξ b).
+  Lemma renaming n m (a b : PTm) (ξ : fin n -> fin m) :
+    rtc ERed.R a b -> rtc ERed.R (ren_PTm a) (ren_PTm b).
   Proof. induction 1; hauto l:on use:ERed.renaming ctrs:rtc. Qed.
 
-  Lemma FromEPar n (a b : PTm n) :
+  Lemma FromEPar n (a b : PTm) :
     EPar.R a b ->
     rtc ERed.R a b.
   Proof.
@@ -1981,18 +1965,18 @@ Module EReds.
       apply ERed.PairEta.
   Qed.
 
-  Lemma FromEPars n (a b : PTm n) :
+  Lemma FromEPars n (a b : PTm) :
     rtc EPar.R a b ->
     rtc ERed.R a b.
   Proof. induction 1; hauto l:on use:FromEPar, @relations.rtc_transitive. Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    rtc ERed.R a b -> rtc ERed.R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    rtc ERed.R a b -> rtc ERed.R (subst_PTm a) (subst_PTm b).
   Proof.
     induction 1; hauto lq:on ctrs:rtc use:ERed.substing.
   Qed.
 
-  Lemma app_inv n (a0 b0 C : PTm n) :
+  Lemma app_inv n (a0 b0 C : PTm) :
     rtc ERed.R (PApp a0 b0) C ->
     exists a1 b1, C = PApp a1 b1 /\
                rtc ERed.R a0 a1 /\
@@ -2006,7 +1990,7 @@ Module EReds.
       hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R.
   Qed.
 
-  Lemma proj_inv n p (a C : PTm n) :
+  Lemma proj_inv n p (a C : PTm) :
     rtc ERed.R (PProj p a) C ->
     exists c, C = PProj p c /\ rtc ERed.R a c.
   Proof.
@@ -2016,7 +2000,7 @@ Module EReds.
       scrush ctrs:rtc,ERed.R inv:ERed.R.
   Qed.
 
-  Lemma bind_inv n p (A : PTm n) B C :
+  Lemma bind_inv n p (A : PTm) B C :
     rtc ERed.R (PBind p A B) C ->
     exists A0 B0, C = PBind p A0 B0 /\ rtc ERed.R A A0 /\ rtc ERed.R B B0.
   Proof.
@@ -2027,7 +2011,7 @@ Module EReds.
     hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
-  Lemma suc_inv n (a : PTm n) C :
+  Lemma suc_inv n (a : PTm) C :
     rtc ERed.R (PSuc a) C ->
     exists b, rtc ERed.R a b /\ C = PSuc b.
   Proof.
@@ -2038,14 +2022,14 @@ Module EReds.
     - hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
-  Lemma zero_inv n (C : PTm n) :
+  Lemma zero_inv n (C : PTm) :
     rtc ERed.R PZero C -> C = PZero.
     move E : PZero => u hu.
     move : E. elim : u C /hu=>//=.
     - hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
-  Lemma ind_inv n P (a : PTm n) b c C :
+  Lemma ind_inv n P (a : PTm) b c C :
     rtc ERed.R (PInd P a b c) C ->
     exists P0 a0 b0 c0,  rtc ERed.R P P0 /\ rtc ERed.R a a0 /\
                       rtc ERed.R b b0 /\ rtc ERed.R c c0 /\
@@ -2063,37 +2047,37 @@ End EReds.
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
 
 Module EJoin.
-  Definition R {n} (a b : PTm n) := exists c, rtc ERed.R a c /\ rtc ERed.R b c.
+  Definition R (a b : PTm) := exists c, rtc ERed.R a c /\ rtc ERed.R b c.
 
-  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm n) :
+  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm) :
     R (PApp a0 b0) (PApp a1 b1) ->
     R a0 a1 /\ R b0 b1.
   Proof.
     hauto lq:on use:EReds.app_inv.
   Qed.
 
-  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm n) :
+  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm) :
     R (PProj p0 a0) (PProj p1 a1) ->
     p0 = p1 /\ R a0 a1.
   Proof.
     hauto lq:on rew:off use:EReds.proj_inv.
   Qed.
 
-  Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
+  Lemma bind_inj n p0 p1 (A0 A1 : PTm) B0 B1 :
     R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ R A0 A1 /\ R B0 B1.
   Proof.
     hauto lq:on rew:off use:EReds.bind_inv.
   Qed.
 
-  Lemma suc_inj n  (A0 A1 : PTm n)  :
+  Lemma suc_inj n  (A0 A1 : PTm)  :
     R (PSuc A0) (PSuc A1) ->
     R A0 A1.
   Proof.
     hauto lq:on rew:off use:EReds.suc_inv.
   Qed.
 
-  Lemma ind_inj n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 :
+  Lemma ind_inj n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
     R P0 P1 /\ R a0 a1 /\ R b0 b1 /\ R c0 c1.
   Proof. hauto q:on use:EReds.ind_inv. Qed.
@@ -2101,7 +2085,7 @@ Module EJoin.
 End EJoin.
 
 Module RERed.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (****************** Beta ***********************)
   | AppAbs a b :
     R (PApp (PAbs a) b)  (subst_PTm (scons b VarPTm) a)
@@ -2162,79 +2146,79 @@ Module RERed.
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
 
-  Lemma ToBetaEta n (a b : PTm n) :
+  Lemma ToBetaEta n (a b : PTm) :
     R a b ->
     ERed.R a b \/ RRed.R a b.
   Proof. induction 1; hauto lq:on db:red. Qed.
 
-  Lemma FromBeta n (a b : PTm n) :
+  Lemma FromBeta n (a b : PTm) :
     RRed.R a b -> RERed.R a b.
   Proof. induction 1; qauto l:on ctrs:R. Qed.
 
-  Lemma FromEta n (a b : PTm n) :
+  Lemma FromEta n (a b : PTm) :
     ERed.R a b -> RERed.R a b.
   Proof. induction 1; qauto l:on ctrs:R. Qed.
 
-  Lemma ToBetaEtaPar n (a b : PTm n) :
+  Lemma ToBetaEtaPar n (a b : PTm) :
     R a b ->
     EPar.R a b \/ RRed.R a b.
   Proof.
     hauto q:on use:ERed.ToEPar, ToBetaEta.
   Qed.
 
-  Lemma sn_preservation n (a b : PTm n) :
+  Lemma sn_preservation n (a b : PTm) :
     R a b ->
     SN a ->
     SN b.
   Proof. hauto q:on use:ToBetaEtaPar, epar_sn_preservation, red_sn_preservation, RPar.FromRRed. Qed.
 
-  Lemma bind_preservation n (a b : PTm n) :
+  Lemma bind_preservation n (a b : PTm) :
     R a b -> isbind a -> isbind b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma univ_preservation n (a b : PTm n) :
+  Lemma univ_preservation n (a b : PTm) :
     R a b -> isuniv a -> isuniv b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma nat_preservation n (a b : PTm n) :
+  Lemma nat_preservation n (a b : PTm) :
     R a b -> isnat a -> isnat b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma sne_preservation n (a b : PTm n) :
+  Lemma sne_preservation n (a b : PTm) :
     R a b -> SNe a -> SNe b.
   Proof.
     hauto q:on use:ToBetaEtaPar, RPar.FromRRed use:red_sn_preservation, epar_sn_preservation.
   Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    RERed.R a b -> RERed.R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    RERed.R a b -> RERed.R (subst_PTm a) (subst_PTm b).
   Proof.
     hauto q:on use:ToBetaEta, FromBeta, FromEta, RRed.substing, ERed.substing.
   Qed.
 
-  Lemma hne_preservation n (a b : PTm n) :
+  Lemma hne_preservation n (a b : PTm) :
     RERed.R a b -> ishne a -> ishne b.
   Proof.  induction 1; sfirstorder. Qed.
 
 End RERed.
 
 Module REReds.
-  Lemma hne_preservation n (a b : PTm n) :
+  Lemma hne_preservation n (a b : PTm) :
     rtc RERed.R a b -> ishne a -> ishne b.
   Proof.  induction 1; eauto using RERed.hne_preservation, rtc_refl, rtc. Qed.
 
-  Lemma sn_preservation n (a b : PTm n) :
+  Lemma sn_preservation n (a b : PTm) :
     rtc RERed.R a b ->
     SN a ->
     SN b.
   Proof. induction 1; eauto using RERed.sn_preservation. Qed.
 
-  Lemma FromRReds n (a b : PTm n) :
+  Lemma FromRReds n (a b : PTm) :
     rtc RRed.R a b ->
     rtc RERed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:RERed.FromBeta. Qed.
 
-  Lemma FromEReds n (a b : PTm n) :
+  Lemma FromEReds n (a b : PTm) :
     rtc ERed.R a b ->
     rtc RERed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:RERed.FromEta. Qed.
@@ -2251,29 +2235,29 @@ Module REReds.
     rtc RERed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R b0 b1 ->
     rtc RERed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R b0 b1 ->
     rtc RERed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm n) :
+  Lemma ProjCong n p (a0 a1  : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm n) :
+  Lemma SucCong n (a0 a1 : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 :
+  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc RERed.R P0 P1 ->
     rtc RERed.R a0 a1 ->
     rtc RERed.R b0 b1 ->
@@ -2281,25 +2265,25 @@ Module REReds.
     rtc RERed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
     rtc RERed.R A0 A1 ->
     rtc RERed.R B0 B1 ->
     rtc RERed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
-  Lemma bind_preservation n (a b : PTm n) :
+  Lemma bind_preservation n (a b : PTm) :
     rtc RERed.R a b -> isbind a -> isbind b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.bind_preservation. Qed.
 
-  Lemma univ_preservation n (a b : PTm n) :
+  Lemma univ_preservation n (a b : PTm) :
     rtc RERed.R a b -> isuniv a -> isuniv b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.univ_preservation. Qed.
 
-  Lemma nat_preservation n (a b : PTm n) :
+  Lemma nat_preservation n (a b : PTm) :
     rtc RERed.R a b -> isnat a -> isnat b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.nat_preservation. Qed.
 
-  Lemma sne_preservation n (a b : PTm n) :
+  Lemma sne_preservation n (a b : PTm) :
     rtc RERed.R a b -> SNe a -> SNe b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.sne_preservation. Qed.
 
@@ -2329,7 +2313,7 @@ Module REReds.
     move : i E. elim : u C /hu; hauto lq:on rew:off inv:RERed.R.
   Qed.
 
-  Lemma hne_app_inv n (a0 b0 C : PTm n) :
+  Lemma hne_app_inv n (a0 b0 C : PTm) :
     rtc RERed.R (PApp a0 b0) C ->
     ishne a0 ->
     exists a1 b1, C = PApp a1 b1 /\
@@ -2344,7 +2328,7 @@ Module REReds.
       hauto lq:on rew:off ctrs:rtc, RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 
-  Lemma hne_proj_inv n p (a C : PTm n) :
+  Lemma hne_proj_inv n p (a C : PTm) :
     rtc RERed.R (PProj p a) C ->
     ishne a ->
     exists c, C = PProj p c /\ rtc RERed.R a c.
@@ -2355,7 +2339,7 @@ Module REReds.
       scrush ctrs:rtc,RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 
-  Lemma hne_ind_inv n P a b c (C : PTm n) :
+  Lemma hne_ind_inv n P a b c (C : PTm) :
     rtc RERed.R (PInd P a b c) C -> ishne a ->
     exists P0 a0 b0 c0, C = PInd P0 a0 b0 c0 /\
                      rtc RERed.R P P0 /\
@@ -2369,47 +2353,47 @@ Module REReds.
       scrush ctrs:rtc,RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    rtc RERed.R a b -> rtc RERed.R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    rtc RERed.R a b -> rtc RERed.R (subst_PTm a) (subst_PTm b).
   Proof.
     induction 1; hauto lq:on ctrs:rtc use:RERed.substing.
   Qed.
 
 
-  Lemma cong_up n m (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma cong_up n m (ρ0 ρ1 : fin n -> PTm) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    (forall i, rtc RERed.R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)).
+    (forall i, rtc RERed.R (up_PTm_PTm0 i) (up_PTm_PTm1 i)).
   Proof. move => h i. destruct i as [i|].
          simpl. rewrite /funcomp.
          substify. by apply substing.
          apply rtc_refl.
   Qed.
 
-  Lemma cong_up2 n m (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma cong_up2 n m (ρ0 ρ1 : fin n -> PTm) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    (forall i, rtc RERed.R (up_PTm_PTm (up_PTm_PTm ρ0) i) (up_PTm_PTm (up_PTm_PTm ρ1) i)).
+    (forall i, rtc RERed.R (up_PTm_PTm (up_PTm_PTm0) i) (up_PTm_PTm (up_PTm_PTm1) i)).
   Proof. hauto l:on use:cong_up. Qed.
 
-  Lemma cong n m (a : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma cong n m (a : PTm) (ρ0 ρ1 : fin n -> PTm) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 a).
+    rtc RERed.R (subst_PTm0 a) (subst_PTm1 a).
   Proof.
     move : m ρ0 ρ1. elim : n / a => /=;
       eauto 20 using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl, IndCong, SucCong, cong_up2.
   Qed.
 
-  Lemma cong' n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
+  Lemma cong' n m (a b : PTm) (ρ0 ρ1 : fin n -> PTm) :
     rtc RERed.R a b ->
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
+    rtc RERed.R (subst_PTm0 a) (subst_PTm1 b).
   Proof.
     move => h0 h1.
-    have : rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 a) by eauto using cong.
+    have : rtc RERed.R (subst_PTm0 a) (subst_PTm1 a) by eauto using cong.
     move => ?. apply : relations.rtc_transitive; eauto.
     hauto l:on use:substing.
   Qed.
 
-  Lemma ToEReds n (a b : PTm n) :
+  Lemma ToEReds n (a b : PTm) :
     nf a ->
     rtc RERed.R a b -> rtc ERed.R a b.
   Proof.
@@ -2417,14 +2401,14 @@ Module REReds.
     induction h; hauto lq:on rew:off ctrs:rtc use:RERed.ToBetaEta, RReds.nf_refl, @rtc_once, ERed.nf_preservation.
   Qed.
 
-  Lemma zero_inv n (C : PTm n) :
+  Lemma zero_inv n (C : PTm) :
     rtc RERed.R PZero C -> C = PZero.
     move E : PZero => u hu.
     move : E. elim : u C /hu=>//=.
     - hauto lq:on rew:off ctrs:rtc, RERed.R inv:RERed.R, rtc.
   Qed.
 
-  Lemma suc_inv n (a : PTm n) C :
+  Lemma suc_inv n (a : PTm) C :
     rtc RERed.R (PSuc a) C ->
     exists b, rtc RERed.R a b /\ C = PSuc b.
   Proof.
@@ -2447,7 +2431,7 @@ Module REReds.
 End REReds.
 
 Module LoRed.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (****************** Beta ***********************)
   | AppAbs a b :
     R (PApp (PAbs a) b)  (subst_PTm (scons b VarPTm) a)
@@ -2515,7 +2499,7 @@ Module LoRed.
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
 
-  Lemma hf_preservation n (a b : PTm n) :
+  Lemma hf_preservation n (a b : PTm) :
     LoRed.R a b ->
     ishf a ->
     ishf b.
@@ -2523,12 +2507,12 @@ Module LoRed.
     move => h. elim : n a b /h=>//=.
   Qed.
 
-  Lemma ToRRed n (a b : PTm n) :
+  Lemma ToRRed n (a b : PTm) :
     LoRed.R a b ->
     RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
 
-  Lemma AppAbs' n a (b : PTm n) u :
+  Lemma AppAbs' n a (b : PTm) u :
     u = (subst_PTm (scons b VarPTm) a) ->
     R (PApp (PAbs a) b)  u.
   Proof. move => ->. apply AppAbs. Qed.
@@ -2538,8 +2522,8 @@ Module LoRed.
     R (@PInd n P (PSuc a) b c) u.
   Proof. move => ->. apply IndSuc. Qed.
 
-  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
+  Lemma renaming n m (a b : PTm) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm a) (ren_PTm b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
@@ -2553,7 +2537,7 @@ Module LoRed.
 End LoRed.
 
 Module LoReds.
-  Lemma hf_preservation n (a b : PTm n) :
+  Lemma hf_preservation n (a b : PTm) :
     rtc LoRed.R a b ->
     ishf a ->
     ishf b.
@@ -2561,7 +2545,7 @@ Module LoReds.
     induction 1; eauto using LoRed.hf_preservation.
   Qed.
 
-  Lemma hf_ne_imp n (a b : PTm n) :
+  Lemma hf_ne_imp n (a b : PTm) :
     rtc LoRed.R a b ->
     ne b ->
     ~~ ishf a.
@@ -2582,34 +2566,34 @@ Module LoReds.
     rtc LoRed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R b0 b1 ->
     ne a1 ->
     rtc LoRed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R b0 b1 ->
     nf a1 ->
     rtc LoRed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm n) :
+  Lemma ProjCong n p (a0 a1  : PTm) :
     rtc LoRed.R a0 a1 ->
     ne a1 ->
     rtc LoRed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
     rtc LoRed.R A0 A1 ->
     rtc LoRed.R B0 B1 ->
     nf A1 ->
     rtc LoRed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 :
+  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R P0 P1 ->
     rtc LoRed.R b0 b1 ->
@@ -2620,7 +2604,7 @@ Module LoReds.
     rtc LoRed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm n) :
+  Lemma SucCong n (a0 a1 : PTm) :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
@@ -2628,9 +2612,9 @@ Module LoReds.
   Local Ltac triv := simpl in *; itauto.
 
   Lemma FromSN_mutual : forall n,
-    (forall (a : PTm n) (_ : SNe a), exists v, rtc LoRed.R a v /\ ne v) /\
-    (forall (a : PTm n) (_ : SN a), exists v, rtc LoRed.R a v /\ nf v) /\
-    (forall (a b : PTm n) (_ : TRedSN a b), LoRed.R a b).
+    (forall (a : PTm) (_ : SNe a), exists v, rtc LoRed.R a v /\ ne v) /\
+    (forall (a : PTm) (_ : SN a), exists v, rtc LoRed.R a v /\ nf v) /\
+    (forall (a b : PTm) (_ : TRedSN a b), LoRed.R a b).
   Proof.
     apply sn_mutual.
     - hauto lq:on ctrs:rtc.
@@ -2666,49 +2650,49 @@ Module LoReds.
   Lemma FromSN : forall n a, @SN n a -> exists v, rtc LoRed.R a v /\ nf v.
   Proof. firstorder using FromSN_mutual. Qed.
 
-  Lemma ToRReds : forall n (a b : PTm n), rtc LoRed.R a b -> rtc RRed.R a b.
+  Lemma ToRReds : forall n (a b : PTm), rtc LoRed.R a b -> rtc RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:LoRed.ToRRed. Qed.
 
 End LoReds.
 
 
-Fixpoint size_PTm {n} (a : PTm n) :=
+Fixpoint size_PTm (a : PTm) :=
   match a with
   | VarPTm _ => 1
-  | PAbs a => 3 + size_PTm a
-  | PApp a b => 1 + Nat.add (size_PTm a) (size_PTm b)
-  | PProj p a => 1 + size_PTm a
-  | PPair a b => 3 + Nat.add (size_PTm a) (size_PTm b)
+  | PAbs a => 3 + size_PTm
+  | PApp a b => 1 + Nat.add (size_PTm) (size_PTm)
+  | PProj p a => 1 + size_PTm
+  | PPair a b => 3 + Nat.add (size_PTm) (size_PTm)
   | PUniv _ => 3
-  | PBind p A B => 3 + Nat.add (size_PTm A) (size_PTm B)
-  | PInd P a b c => 3 + size_PTm P + size_PTm a + size_PTm b + size_PTm c
+  | PBind p A B => 3 + Nat.add (size_PTm) (size_PTm)
+  | PInd P a b c => 3 + size_PTm + size_PTm + size_PTm + size_PTm
   | PNat => 3
-  | PSuc a => 3 + size_PTm a
+  | PSuc a => 3 + size_PTm
   | PZero => 3
   | PBot => 1
   end.
 
-Lemma size_PTm_ren n m (ξ : fin n -> fin m) a : size_PTm (ren_PTm ξ a) = size_PTm a.
+Lemma size_PTm_ren n m (ξ : fin n -> fin m) a : size_PTm (ren_PTm a) = size_PTm.
 Proof.
   move : m ξ. elim : n / a => //=; scongruence.
 Qed.
 
 #[export]Hint Rewrite size_PTm_ren : sizetm.
 
-Lemma ered_size {n} (a b : PTm n) :
+Lemma ered_size (a b : PTm) :
   ERed.R a b ->
-  size_PTm b < size_PTm a.
+  size_PTm < size_PTm.
 Proof.
   move => h. elim : n a b /h; hauto l:on rew:db:sizetm.
 Qed.
 
-Lemma ered_sn n (a : PTm n) : sn ERed.R a.
+Lemma ered_sn n (a : PTm) : sn ERed.R a.
 Proof.
   hauto lq:on rew:off use:size_PTm_ren, ered_size,
           well_founded_lt_compat unfold:well_founded.
 Qed.
 
-Lemma ered_local_confluence n (a b c : PTm n) :
+Lemma ered_local_confluence n (a b c : PTm) :
   ERed.R a b ->
   ERed.R a c ->
   exists d, rtc ERed.R b d  /\ rtc ERed.R c d.
@@ -2814,7 +2798,7 @@ Proof.
   - qauto l:on inv:ERed.R ctrs:rtc use:EReds.SucCong.
 Qed.
 
-Lemma ered_confluence n (a b c : PTm n) :
+Lemma ered_confluence n (a b c : PTm) :
   rtc ERed.R a b ->
   rtc ERed.R a c ->
   exists d, rtc ERed.R b d  /\ rtc ERed.R c d.
@@ -2822,18 +2806,18 @@ Proof.
   sfirstorder use:relations.locally_confluent_confluent, ered_sn, ered_local_confluence.
 Qed.
 
-Lemma red_confluence n (a b c : PTm n) :
+Lemma red_confluence n (a b c : PTm) :
   rtc RRed.R a b ->
   rtc RRed.R a c ->
   exists d, rtc RRed.R b d  /\ rtc RRed.R c d.
-  suff : rtc RPar.R a b -> rtc RPar.R a c -> exists d : PTm n, rtc RPar.R b d /\ rtc RPar.R c d
+  suff : rtc RPar.R a b -> rtc RPar.R a c -> exists d : PTm, rtc RPar.R b d /\ rtc RPar.R c d
              by hauto lq:on use:RReds.RParIff.
   apply relations.diamond_confluent.
   rewrite /relations.diamond.
   eauto using RPar.diamond.
 Qed.
 
-Lemma red_uniquenf n (a b c : PTm n) :
+Lemma red_uniquenf n (a b c : PTm) :
   rtc RRed.R a b ->
   rtc RRed.R a c ->
   nf b ->
@@ -2848,20 +2832,20 @@ Proof.
 Qed.
 
 Module NeEPars.
-  Lemma R_nonelim_nf n (a b : PTm n) :
+  Lemma R_nonelim_nf n (a b : PTm) :
     rtc NeEPar.R_nonelim a b ->
     nf b ->
     nf a.
   Proof. induction 1; sfirstorder use:NeEPar.R_elim_nf. Qed.
 
-  Lemma ToEReds : forall n, (forall (a b : PTm n), rtc NeEPar.R_nonelim a b -> rtc ERed.R a b).
+  Lemma ToEReds : forall n, (forall (a b : PTm), rtc NeEPar.R_nonelim a b -> rtc ERed.R a b).
   Proof.
     induction 1; hauto l:on use:NeEPar.ToEPar, EReds.FromEPar, @relations.rtc_transitive.
   Qed.
 End NeEPars.
 
 
-Lemma rered_standardization n (a c : PTm n) :
+Lemma rered_standardization n (a c : PTm) :
   SN a ->
   rtc RERed.R a c ->
   exists b, rtc RRed.R a b /\ rtc NeEPar.R_nonelim b c.
@@ -2878,7 +2862,7 @@ Proof.
   - hauto lq:on ctrs:rtc use:red_sn_preservation, RPar.FromRRed.
 Qed.
 
-Lemma rered_confluence n (a b c : PTm n) :
+Lemma rered_confluence n (a b c : PTm) :
   SN a ->
   rtc RERed.R a b ->
   rtc RERed.R a c ->
@@ -2912,14 +2896,14 @@ Qed.
 
 (* Beta joinability *)
 Module BJoin.
-  Definition R {n} (a b : PTm n) := exists c, rtc RRed.R a c /\ rtc RRed.R b c.
-  Lemma refl n (a : PTm n) : R a a.
+  Definition R (a b : PTm) := exists c, rtc RRed.R a c /\ rtc RRed.R b c.
+  Lemma refl n (a : PTm) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
-  Lemma symmetric n (a b : PTm n) : R a b -> R b a.
+  Lemma symmetric n (a b : PTm) : R a b -> R b a.
   Proof. sfirstorder unfold:R. Qed.
 
-  Lemma transitive n (a b c : PTm n) : R a b -> R b c -> R a c.
+  Lemma transitive n (a b c : PTm) : R a b -> R b c -> R a c.
   Proof.
     rewrite /R.
     move => [ab [ha +]] [bc [+ hc]].
@@ -2933,7 +2917,7 @@ Module BJoin.
   (*   R (PAbs a) (PAbs b). *)
   (* Proof. hauto lq:on use:RReds.AbsCong unfold:R. Qed. *)
 
-  (* Lemma AppCong n (a0 a1 b0 b1 : PTm n) : *)
+  (* Lemma AppCong n (a0 a1 b0 b1 : PTm) : *)
   (*   R a0 a1 -> *)
   (*   R b0 b1 -> *)
   (*   R (PApp a0 b0) (PApp a1 b1). *)
@@ -2941,21 +2925,21 @@ Module BJoin.
 End BJoin.
 
 Module DJoin.
-  Definition R {n} (a b : PTm n) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.
+  Definition R (a b : PTm) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.
 
-  Lemma refl n (a : PTm n) : R a a.
+  Lemma refl n (a : PTm) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
-  Lemma FromEJoin n (a b : PTm n) : EJoin.R a b -> DJoin.R a b.
+  Lemma FromEJoin n (a b : PTm) : EJoin.R a b -> DJoin.R a b.
   Proof. hauto q:on use:REReds.FromEReds. Qed.
 
-  Lemma ToEJoin n (a b : PTm n) : nf a -> nf b -> DJoin.R a b -> EJoin.R a b.
+  Lemma ToEJoin n (a b : PTm) : nf a -> nf b -> DJoin.R a b -> EJoin.R a b.
   Proof. hauto q:on use:REReds.ToEReds. Qed.
 
-  Lemma symmetric n (a b : PTm n) : R a b -> R b a.
+  Lemma symmetric n (a b : PTm) : R a b -> R b a.
   Proof. sfirstorder unfold:R. Qed.
 
-  Lemma transitive n (a b c : PTm n) : SN b -> R a b -> R b c -> R a c.
+  Lemma transitive n (a b c : PTm) : SN b -> R a b -> R b c -> R a c.
   Proof.
     rewrite /R.
     move => + [ab [ha +]] [bc [+ hc]].
@@ -2969,30 +2953,30 @@ Module DJoin.
     R (PAbs a) (PAbs b).
   Proof. hauto lq:on use:REReds.AbsCong unfold:R. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
     R a0 a1 ->
     R b0 b1 ->
     R (PApp a0 b0) (PApp a1 b1).
   Proof. hauto lq:on use:REReds.AppCong unfold:R. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
+  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
     R a0 a1 ->
     R b0 b1 ->
     R (PPair a0 b0) (PPair a1 b1).
   Proof. hauto q:on use:REReds.PairCong. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm n) :
+  Lemma ProjCong n p (a0 a1  : PTm) :
     R a0 a1 ->
     R (PProj p a0) (PProj p a1).
   Proof. hauto q:on use:REReds.ProjCong. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
     R A0 A1 ->
     R B0 B1 ->
     R (PBind p A0 B0) (PBind p A1 B1).
   Proof. hauto q:on use:REReds.BindCong. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 :
+  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     R P0 P1 ->
     R a0 a1 ->
     R b0 b1 ->
@@ -3000,12 +2984,12 @@ Module DJoin.
     R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. hauto q:on use:REReds.IndCong. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm n) :
+  Lemma SucCong n (a0 a1 : PTm) :
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
   Proof. qauto l:on use:REReds.SucCong. Qed.
 
-  Lemma FromRedSNs n (a b : PTm n) :
+  Lemma FromRedSNs n (a b : PTm) :
     rtc TRedSN a b ->
     R a b.
   Proof.
@@ -3013,7 +2997,7 @@ Module DJoin.
     sfirstorder use:@rtc_refl unfold:R.
   Qed.
 
-  Lemma sne_nat_noconf n (a b : PTm n) :
+  Lemma sne_nat_noconf n (a b : PTm) :
     R a b -> SNe a -> isnat b -> False.
   Proof.
     move => [c [? ?]] *.
@@ -3021,7 +3005,7 @@ Module DJoin.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma sne_bind_noconf n (a b : PTm n) :
+  Lemma sne_bind_noconf n (a b : PTm) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
     move => [c [? ?]] *.
@@ -3029,14 +3013,14 @@ Module DJoin.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma sne_univ_noconf n (a b : PTm n) :
+  Lemma sne_univ_noconf n (a b : PTm) :
     R a b -> SNe a -> isuniv b -> False.
   Proof.
     hauto q:on use:REReds.sne_preservation,
           REReds.univ_preservation inv:SNe.
   Qed.
 
-  Lemma bind_univ_noconf n (a b : PTm n) :
+  Lemma bind_univ_noconf n (a b : PTm) :
     R a b -> isbind a -> isuniv b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3046,7 +3030,7 @@ Module DJoin.
     case : c => //=; itauto.
   Qed.
 
-  Lemma hne_univ_noconf n (a b : PTm n) :
+  Lemma hne_univ_noconf n (a b : PTm) :
     R a b -> ishne a -> isuniv b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3057,7 +3041,7 @@ Module DJoin.
     case : c => //=.
   Qed.
 
-  Lemma hne_bind_noconf n (a b : PTm n) :
+  Lemma hne_bind_noconf n (a b : PTm) :
     R a b -> ishne a -> isbind b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3068,7 +3052,7 @@ Module DJoin.
     case : c => //=.
   Qed.
 
-  Lemma hne_nat_noconf n (a b : PTm n) :
+  Lemma hne_nat_noconf n (a b : PTm) :
     R a b -> ishne a -> isnat b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3076,7 +3060,7 @@ Module DJoin.
     clear. case : c => //=; itauto.
   Qed.
 
-  Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
+  Lemma bind_inj n p0 p1 (A0 A1 : PTm) B0 B1 :
     DJoin.R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ DJoin.R A0 A1 /\ DJoin.R B0 B1.
   Proof.
@@ -3094,14 +3078,14 @@ Module DJoin.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
-  Lemma suc_inj n  (A0 A1 : PTm n)  :
+  Lemma suc_inj n  (A0 A1 : PTm)  :
     R (PSuc A0) (PSuc A1) ->
     R A0 A1.
   Proof.
     hauto lq:on rew:off use:REReds.suc_inv.
   Qed.
 
-  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm n) :
+  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm) :
     R (PApp a0 b0) (PApp a1 b1) ->
     ishne a0 ->
     ishne a1 ->
@@ -3110,7 +3094,7 @@ Module DJoin.
     hauto lq:on use:REReds.hne_app_inv.
   Qed.
 
-  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm n) :
+  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm) :
     R (PProj p0 a0) (PProj p1 a1) ->
     ishne a0 ->
     ishne a1 ->
@@ -3119,45 +3103,45 @@ Module DJoin.
     sauto lq:on use:REReds.hne_proj_inv.
   Qed.
 
-  Lemma FromRRed0 n (a b : PTm n) :
+  Lemma FromRRed0 n (a b : PTm) :
     RRed.R a b -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
   Qed.
 
-  Lemma FromRRed1 n (a b : PTm n) :
+  Lemma FromRRed1 n (a b : PTm) :
     RRed.R b a -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
   Qed.
 
-  Lemma FromRReds n (a b : PTm n) :
+  Lemma FromRReds n (a b : PTm) :
     rtc RRed.R b a -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:REReds.FromRReds unfold:R.
   Qed.
 
-  Lemma FromBJoin n (a b : PTm n) :
+  Lemma FromBJoin n (a b : PTm) :
     BJoin.R a b -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:REReds.FromRReds unfold:R.
   Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    R a b -> R (subst_PTm a) (subst_PTm b).
   Proof.
     hauto lq:on rew:off ctrs:rtc unfold:R use:REReds.substing.
   Qed.
 
-  Lemma renaming n m (a b : PTm n) (ρ : fin n -> fin m) :
-    R a b -> R (ren_PTm ρ a) (ren_PTm ρ b).
+  Lemma renaming n m (a b : PTm) (ρ : fin n -> fin m) :
+    R a b -> R (ren_PTm a) (ren_PTm b).
   Proof. substify. apply substing. Qed.
 
-  Lemma weakening n (a b : PTm n)  :
+  Lemma weakening n (a b : PTm)  :
     R a b -> R (ren_PTm shift a) (ren_PTm shift b).
   Proof. apply renaming. Qed.
 
-  Lemma cong n m (a : PTm (S n)) c d (ρ : fin n -> PTm m) :
+  Lemma cong n m (a : PTm (S n)) c d (ρ : fin n -> PTm) :
     R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) a).
   Proof.
     rewrite /R. move => [cd [h0 h1]].
@@ -3165,7 +3149,7 @@ Module DJoin.
     hauto q:on ctrs:rtc inv:option use:REReds.cong.
   Qed.
 
-  Lemma cong' n m (a b : PTm (S n)) c d (ρ : fin n -> PTm m) :
+  Lemma cong' n m (a b : PTm (S n)) c d (ρ : fin n -> PTm) :
     R a b ->
     R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) b).
   Proof.
@@ -3174,7 +3158,7 @@ Module DJoin.
     hauto q:on ctrs:rtc inv:option use:REReds.cong'.
   Qed.
 
-  Lemma pair_inj n (a0 a1 b0 b1 : PTm n) :
+  Lemma pair_inj n (a0 a1 b0 b1 : PTm) :
     SN (PPair a0 b0) ->
     SN (PPair a1 b1) ->
     R (PPair a0 b0) (PPair a1 b1) ->
@@ -3201,7 +3185,7 @@ Module DJoin.
     split; eauto using transitive.
   Qed.
 
-  Lemma ejoin_pair_inj n (a0 a1 b0 b1 : PTm n) :
+  Lemma ejoin_pair_inj n (a0 a1 b0 b1 : PTm) :
     nf a0 -> nf b0 -> nf a1 -> nf b1 ->
     EJoin.R (PPair a0 b0) (PPair a1 b1) ->
     EJoin.R a0 a1 /\ EJoin.R b0 b1.
@@ -3250,7 +3234,7 @@ Module DJoin.
     hauto l:on use:ToEJoin.
   Qed.
 
-  Lemma standardization n (a b : PTm n) :
+  Lemma standardization n (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc RRed.R a va /\ rtc RRed.R b vb /\ nf va /\ nf vb /\ EJoin.R va vb.
   Proof.
@@ -3270,7 +3254,7 @@ Module DJoin.
     hauto q:on use:NeEPars.ToEReds unfold:EJoin.R.
   Qed.
 
-  Lemma standardization_lo n (a b : PTm n) :
+  Lemma standardization_lo n (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc LoRed.R a va /\ rtc LoRed.R b vb /\ nf va /\ nf vb /\ EJoin.R va vb.
   Proof.
@@ -3288,7 +3272,7 @@ End DJoin.
 
 
 Module Sub1.
-  Inductive R {n} : PTm n -> PTm n -> Prop :=
+  Inductive R : PTm -> PTm -> Prop :=
   | Refl a :
     R a a
   | Univ i j :
@@ -3303,21 +3287,21 @@ Module Sub1.
     R B0 B1 ->
     R (PBind PSig A0 B0) (PBind PSig A1 B1).
 
-  Lemma transitive0 n (A B C : PTm n) :
+  Lemma transitive0 n (A B C : PTm) :
     R A B -> (R B C -> R A C) /\ (R C A -> R C B).
   Proof.
     move => h. move : C.
     elim : n A B /h; hauto lq:on ctrs:R inv:R solve+:lia.
   Qed.
 
-  Lemma transitive n (A B C : PTm n) :
+  Lemma transitive n (A B C : PTm) :
     R A B -> R B C -> R A C.
   Proof. hauto q:on use:transitive0. Qed.
 
-  Lemma refl n (A : PTm n) : R A A.
+  Lemma refl n (A : PTm) : R A A.
   Proof. sfirstorder. Qed.
 
-  Lemma commutativity0 n (A B C : PTm n) :
+  Lemma commutativity0 n (A B C : PTm) :
     R A B ->
     (RERed.R A C ->
     exists D, RERed.R B D /\ R C D) /\
@@ -3328,7 +3312,7 @@ Module Sub1.
     elim : n A B / h; hauto lq:on ctrs:RERed.R, R inv:RERed.R.
   Qed.
 
-  Lemma commutativity_Ls n (A B C : PTm n) :
+  Lemma commutativity_Ls n (A B C : PTm) :
     R A B ->
     rtc RERed.R A C ->
     exists D, rtc RERed.R B D /\ R C D.
@@ -3336,7 +3320,7 @@ Module Sub1.
     move => + h. move : B. induction h; sauto lq:on use:commutativity0.
   Qed.
 
-  Lemma commutativity_Rs n (A B C : PTm n) :
+  Lemma commutativity_Rs n (A B C : PTm) :
     R A B ->
     rtc RERed.R B C ->
     exists D, rtc RERed.R A D /\ R D C.
@@ -3345,46 +3329,46 @@ Module Sub1.
   Qed.
 
   Lemma sn_preservation  : forall n,
-  (forall (a : PTm n) (s : SNe a), forall b, R a b \/ R b a -> a = b) /\
-  (forall (a : PTm n) (s : SN a), forall b, R a b \/ R b a -> SN b) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall c, R a c \/ R c a -> a = c).
+  (forall (a : PTm) (s : SNe a), forall b, R a b \/ R b a -> a = b) /\
+  (forall (a : PTm) (s : SN a), forall b, R a b \/ R b a -> SN b) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall c, R a c \/ R c a -> a = c).
   Proof.
     apply sn_mutual; hauto lq:on inv:R ctrs:SN.
   Qed.
 
-  Lemma bind_preservation n (a b : PTm n) :
+  Lemma bind_preservation n (a b : PTm) :
     R a b -> isbind a = isbind b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma univ_preservation n (a b : PTm n) :
+  Lemma univ_preservation n (a b : PTm) :
     R a b -> isuniv a = isuniv b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma sne_preservation n (a b : PTm n) :
+  Lemma sne_preservation n (a b : PTm) :
     R a b -> SNe a <-> SNe b.
   Proof. hfcrush use:sn_preservation. Qed.
 
-  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
+  Lemma renaming n m (a b : PTm) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm a) (ren_PTm b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h; hauto lq:on ctrs:R.
   Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    R a b -> R (subst_PTm a) (subst_PTm b).
   Proof. move => h. move : m ρ. elim : n a b /h; hauto lq:on ctrs:R. Qed.
 
-  Lemma hne_refl n (a b : PTm n) :
+  Lemma hne_refl n (a b : PTm) :
     ishne a \/ ishne b -> R a b -> a = b.
   Proof. hauto q:on inv:R. Qed.
 
 End Sub1.
 
 Module ESub.
-  Definition R {n} (a b : PTm n) := exists c0 c1, rtc ERed.R a c0 /\ rtc ERed.R b c1 /\ Sub1.R c0 c1.
+  Definition R (a b : PTm) := exists c0 c1, rtc ERed.R a c0 /\ rtc ERed.R b c1 /\ Sub1.R c0 c1.
 
-  Lemma pi_inj n (A0 A1 : PTm n) B0 B1 :
+  Lemma pi_inj n (A0 A1 : PTm) B0 B1 :
     R (PBind PPi A0 B0) (PBind PPi A1 B1) ->
     R A1 A0 /\ R B0 B1.
   Proof.
@@ -3394,7 +3378,7 @@ Module ESub.
     sauto lq:on rew:off inv:Sub1.R.
   Qed.
 
-  Lemma sig_inj n (A0 A1 : PTm n) B0 B1 :
+  Lemma sig_inj n (A0 A1 : PTm) B0 B1 :
     R (PBind PSig A0 B0) (PBind PSig A1 B1) ->
     R A0 A1 /\ R B0 B1.
   Proof.
@@ -3404,7 +3388,7 @@ Module ESub.
     sauto lq:on rew:off inv:Sub1.R.
   Qed.
 
-  Lemma suc_inj n (a b : PTm n) :
+  Lemma suc_inj n (a b : PTm) :
     R (PSuc a) (PSuc b) ->
     R a b.
   Proof.
@@ -3414,12 +3398,12 @@ Module ESub.
 End ESub.
 
 Module Sub.
-  Definition R {n} (a b : PTm n) := exists c d, rtc RERed.R a c /\ rtc RERed.R b d /\ Sub1.R c d.
+  Definition R (a b : PTm) := exists c d, rtc RERed.R a c /\ rtc RERed.R b d /\ Sub1.R c d.
 
-  Lemma refl n (a : PTm n) : R a a.
+  Lemma refl n (a : PTm) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
-  Lemma ToJoin n (a b : PTm n) :
+  Lemma ToJoin n (a b : PTm) :
     ishne a \/ ishne b ->
     R a b ->
     DJoin.R a b.
@@ -3429,7 +3413,7 @@ Module Sub.
     hauto lq:on rew:off use:Sub1.hne_refl.
   Qed.
 
-  Lemma transitive n (a b c : PTm n) : SN b -> R a b -> R b c -> R a c.
+  Lemma transitive n (a b c : PTm) : SN b -> R a b -> R b c -> R a c.
   Proof.
     rewrite /R.
     move => h  [a0][b0][ha][hb]ha0b0 [b1][c0][hb'][hc]hb1c0.
@@ -3441,10 +3425,10 @@ Module Sub.
     exists a',c'; hauto l:on use:Sub1.transitive, @relations.rtc_transitive.
   Qed.
 
-  Lemma FromJoin n (a b : PTm n) : DJoin.R a b -> R a b.
+  Lemma FromJoin n (a b : PTm) : DJoin.R a b -> R a b.
   Proof. sfirstorder. Qed.
 
-  Lemma PiCong n (A0 A1 : PTm n) B0 B1 :
+  Lemma PiCong n (A0 A1 : PTm) B0 B1 :
     R A1 A0 ->
     R B0 B1 ->
     R (PBind PPi A0 B0) (PBind PPi A1 B1).
@@ -3456,7 +3440,7 @@ Module Sub.
     repeat split; eauto using REReds.BindCong, Sub1.PiCong.
   Qed.
 
-  Lemma SigCong n (A0 A1 : PTm n) B0 B1 :
+  Lemma SigCong n (A0 A1 : PTm) B0 B1 :
     R A0 A1 ->
     R B0 B1 ->
     R (PBind PSig A0 B0) (PBind PSig A1 B1).
@@ -3473,7 +3457,7 @@ Module Sub.
     @R n (PUniv i) (PUniv j).
   Proof. hauto lq:on ctrs:Sub1.R, rtc. Qed.
 
-  Lemma sne_nat_noconf n (a b : PTm n) :
+  Lemma sne_nat_noconf n (a b : PTm) :
     R a b -> SNe a -> isnat b -> False.
   Proof.
     move => [c [d [h0 [h1 h2]]]] *.
@@ -3481,7 +3465,7 @@ Module Sub.
     move : h2. clear. hauto q:on inv:Sub1.R, SNe.
   Qed.
 
-  Lemma nat_sne_noconf n (a b : PTm n) :
+  Lemma nat_sne_noconf n (a b : PTm) :
     R a b -> isnat a -> SNe b -> False.
   Proof.
     move => [c [d [h0 [h1 h2]]]] *.
@@ -3489,7 +3473,7 @@ Module Sub.
     move : h2. clear. hauto q:on inv:Sub1.R, SNe.
   Qed.
 
-  Lemma sne_bind_noconf n (a b : PTm n) :
+  Lemma sne_bind_noconf n (a b : PTm) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
     rewrite /R.
@@ -3499,7 +3483,7 @@ Module Sub.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma hne_bind_noconf n (a b : PTm n) :
+  Lemma hne_bind_noconf n (a b : PTm) :
     R a b -> ishne a -> isbind b -> False.
   Proof.
     rewrite /R.
@@ -3510,7 +3494,7 @@ Module Sub.
     clear. case : d => //=.
   Qed.
 
-  Lemma bind_sne_noconf n (a b : PTm n) :
+  Lemma bind_sne_noconf n (a b : PTm) :
     R a b -> SNe b -> isbind a -> False.
   Proof.
     rewrite /R.
@@ -3520,14 +3504,14 @@ Module Sub.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma sne_univ_noconf n (a b : PTm n) :
+  Lemma sne_univ_noconf n (a b : PTm) :
     R a b -> SNe a -> isuniv b -> False.
   Proof.
     hauto l:on use:REReds.sne_preservation,
           REReds.univ_preservation, Sub1.sne_preservation, Sub1.univ_preservation inv:SNe.
   Qed.
 
-  Lemma univ_sne_noconf n (a b : PTm n) :
+  Lemma univ_sne_noconf n (a b : PTm) :
     R a b -> SNe b -> isuniv a -> False.
   Proof.
     move => [c[d] [? []]] *.
@@ -3537,7 +3521,7 @@ Module Sub.
     clear. case : c => //=. inversion 2.
   Qed.
 
-  Lemma univ_nat_noconf n (a b : PTm n)  :
+  Lemma univ_nat_noconf n (a b : PTm)  :
     R a b -> isuniv b -> isnat a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3547,7 +3531,7 @@ Module Sub.
     clear. case : d => //=.
   Qed.
 
-  Lemma nat_univ_noconf n (a b : PTm n)  :
+  Lemma nat_univ_noconf n (a b : PTm)  :
     R a b -> isnat b -> isuniv a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3557,7 +3541,7 @@ Module Sub.
     clear. case : d => //=.
   Qed.
 
-  Lemma bind_nat_noconf n (a b : PTm n)  :
+  Lemma bind_nat_noconf n (a b : PTm)  :
     R a b -> isbind b -> isnat a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3568,7 +3552,7 @@ Module Sub.
     case : d h1 => //=.
   Qed.
 
-  Lemma nat_bind_noconf n (a b : PTm n)  :
+  Lemma nat_bind_noconf n (a b : PTm)  :
     R a b -> isnat b -> isbind a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3579,7 +3563,7 @@ Module Sub.
     case : d h1 => //=.
   Qed.
 
-  Lemma bind_univ_noconf n (a b : PTm n) :
+  Lemma bind_univ_noconf n (a b : PTm) :
     R a b -> isbind a -> isuniv b -> False.
   Proof.
     move => [c[d] [h0 [h1 h1']]] h2 h3.
@@ -3590,7 +3574,7 @@ Module Sub.
     case : c => //=; itauto.
   Qed.
 
-  Lemma univ_bind_noconf n (a b : PTm n) :
+  Lemma univ_bind_noconf n (a b : PTm) :
     R a b -> isbind b -> isuniv a -> False.
   Proof.
     move => [c[d] [h0 [h1 h1']]] h2 h3.
@@ -3601,7 +3585,7 @@ Module Sub.
     case : c => //=; itauto.
   Qed.
 
-  Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
+  Lemma bind_inj n p0 p1 (A0 A1 : PTm) B0 B1 :
     R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ (if p0 is PPi then R A1 A0 else R A0 A1) /\ R B0 B1.
   Proof.
@@ -3618,7 +3602,7 @@ Module Sub.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
-  Lemma suc_inj n  (A0 A1 : PTm n)  :
+  Lemma suc_inj n  (A0 A1 : PTm)  :
     R (PSuc A0) (PSuc A1) ->
     R A0 A1.
   Proof.
@@ -3626,7 +3610,7 @@ Module Sub.
   Qed.
 
 
-  Lemma cong n m (a b : PTm (S n)) c d (ρ : fin n -> PTm m) :
+  Lemma cong n m (a b : PTm (S n)) c d (ρ : fin n -> PTm) :
     R a b -> DJoin.R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) b).
   Proof.
     rewrite /R.
@@ -3639,18 +3623,18 @@ Module Sub.
     eauto using Sub1.substing.
   Qed.
 
-  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
-    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
+    R a b -> R (subst_PTm a) (subst_PTm b).
   Proof.
     rewrite /R.
     move => [a0][b0][h0][h1]h2.
     hauto ctrs:rtc use:REReds.cong', Sub1.substing.
   Qed.
 
-  Lemma ToESub n (a b : PTm n) : nf a -> nf b -> R a b -> ESub.R a b.
+  Lemma ToESub n (a b : PTm) : nf a -> nf b -> R a b -> ESub.R a b.
   Proof. hauto q:on use:REReds.ToEReds. Qed.
 
-  Lemma standardization n (a b : PTm n) :
+  Lemma standardization n (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc RRed.R a va /\ rtc RRed.R b vb /\ nf va /\ nf vb /\ ESub.R va vb.
   Proof.
@@ -3668,7 +3652,7 @@ Module Sub.
     hauto lq:on.
   Qed.
 
-  Lemma standardization_lo n (a b : PTm n) :
+  Lemma standardization_lo n (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc LoRed.R a va /\ rtc LoRed.R b vb /\ nf va /\ nf vb /\ ESub.R va vb.
   Proof.

From 551dd5c17ca15a29d6bf1a45b5204463ba888f48 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 2 Mar 2025 20:19:11 -0500
Subject: [PATCH 03/24] Fix the fv proof

---
 theories/fp_red.v | 178 ++++++++++++++++++++++++++--------------------
 1 file changed, 101 insertions(+), 77 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5f7ea0d..1e7863b 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1788,109 +1788,133 @@ Module ERed.
   Qed.
 
   (* Characterization of variable freeness conditions *)
-  Definition tm_i_free a (i : fin n) := exists m (ξ ξ0 : fin n -> fin m), ξ i <> ξ0 i /\ ren_PTm a = ren_PTm0 a.
+  Definition tm_i_free a (i : nat) := exists (ξ ξ0 : nat -> nat), ξ i <> ξ0 i /\ ren_PTm ξ a = ren_PTm ξ0 a.
 
-  Lemma subst_differ_one_ren_up n m i (ξ0 ξ1 : fin n -> fin m) :
+  Lemma subst_differ_one_ren_up i (ξ0 ξ1 : nat -> nat) :
     (forall j, i <> j -> ξ0 j = ξ1 j) ->
-    (forall j, shift i <> j -> upRen_PTm_PTm0 j =  upRen_PTm_PTm1 j).
+    (forall j, shift i <> j -> upRen_PTm_PTm ξ0 j =  upRen_PTm_PTm ξ1 j).
   Proof.
     move => hξ.
-    destruct j. asimpl.
-    move => h. have :  i<> f by hauto lq:on rew:off inv:option.
+    case => // j.
+    asimpl.
+    move => h. have :  i<> j by hauto lq:on rew:off.
     move /hξ. by rewrite /funcomp => ->.
-    done.
   Qed.
 
-  Lemma tm_free_ren_any n a i :
+  Lemma tm_free_ren_any a i :
     tm_i_free a i ->
-    forall m (ξ0 ξ1 : fin n -> fin m), (forall j, i <> j -> ξ0 j = ξ1 j) ->
-             ren_PTm0 a = ren_PTm1 a.
+    forall (ξ0 ξ1 : nat -> nat), (forall j, i <> j -> ξ0 j = ξ1 j) ->
+             ren_PTm ξ0 a = ren_PTm ξ1 a.
   Proof.
     rewrite /tm_i_free.
     move => [+ [+ [+ +]]].
     move : i.
-    elim : n / a => n.
+    elim : a.
     - hauto q:on.
-    - move => a iha i m ρ0 ρ1 [] => h [] h' m' ξ0 ξ1 hξ /=.
+    - move => a iha i ρ0 ρ1 h [] h' ξ0 ξ1 hξ /=.
       f_equal. move /subst_differ_one_ren_up in hξ.
       move /(_ (shift i)) in iha.
       move : iha hξ; move/[apply].
-      apply=>//. split; eauto.
-      asimpl. rewrite /funcomp. hauto l:on.
+      apply; cycle 1; eauto.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
-    - move => p A ihA a iha i m ρ0 ρ1 [] ? h m' ξ0 ξ1 hξ /=.
+    - move => p A ihA a iha i ρ0 ρ1 hρ [] ? h ξ0 ξ1 hξ /=.
       f_equal. hauto lq:on rew:off.
       move /subst_differ_one_ren_up in hξ.
       move /(_ (shift i)) in iha.
-      move : iha hξ. repeat move/[apply].
-      move /(_ _ (upRen_PTm_PTm0) (upRen_PTm_PTm1)).
-      apply. hauto l:on.
+      move : iha (hξ). repeat move/[apply].
+      move /(_ (upRen_PTm_PTm ρ0) (upRen_PTm_PTm ρ1)).
+      apply. simpl. rewrite /funcomp. scongruence.
+      sfirstorder.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
-    - hauto lq:on rew:off.
-    - admit.
-  Admitted.
+    - move => P ihP a iha b ihb c ihc i ρ0 ρ1 hρ /= []hP
+               ha hb hc ξ0 ξ1 hξ.
+      f_equal;eauto.
+      apply : ihP; cycle 1; eauto using subst_differ_one_ren_up.
+      apply : ihc; cycle 1; eauto using subst_differ_one_ren_up.
+      hauto l:on.
+  Qed.
 
-  Lemma antirenaming n m (a : PTm) (b : PTm) (ξ : fin n -> fin m) :
+  Lemma antirenaming (a : PTm) (b : PTm) (ξ : nat -> nat) :
     (forall i j, ξ i = ξ j -> i = j) ->
-    R (ren_PTm a) b -> exists b0, R a b0 /\ ren_PTm b0 = b.
+    R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
   Proof.
     move => hξ.
-    move E : (ren_PTm a) => u hu.
-    move : n ξ a hξ E.
-    elim : m u b / hu; try solve_anti_ren.
-    - move => n a m ξ []//=.
+    move E : (ren_PTm ξ a) => u hu.
+    move : ξ a hξ E.
+    elim : u b / hu; try solve_anti_ren.
+    - move => a ξ []//=.
       move => u hξ [].
       case : u => //=.
       move => u0 u1 [].
       case : u1 => //=.
       move => i /[swap] [].
       case : i => //= _ h.
-      have : exists p, ren_PTm shift p = u0 by admit.
+      suff : exists p, ren_PTm shift p = u0.
       move => [p ?]. subst.
       move : h. asimpl.
-      replace (ren_PTm (funcomp shift ξ) p) with
-        (ren_PTm shift (ren_PTm p)); last by asimpl.
+      replace (ren_PTm (funcomp S ξ) p) with
+        (ren_PTm shift (ren_PTm ξ p)); last by asimpl.
       move /ren_injective.
-      move /(_ ltac:(hauto l:on)).
+      move /(_ ltac:(hauto l:on unfold:ren_inj)).
       move => ?. subst.
       exists p. split=>//. apply AppEta.
-    - move => n a m ξ [] //=.
+      set u := ren_PTm (scons 0 id) u0.
+      suff : ren_PTm shift u = u0 by eauto.
+      subst u.
+      asimpl.
+      have hE : u0 = ren_PTm id u0 by asimpl.
+      rewrite {2}hE. move{hE}.
+      apply (tm_free_ren_any _ 0); last by qauto l:on inv:nat.
+      rewrite /tm_i_free.
+      have h' := h.
+      apply f_equal with (f := ren_PTm (scons 0 id)) in h.
+      apply f_equal with (f := ren_PTm (scons 1 id)) in h'.
+      move : h'. asimpl => *. subst.
+      move : h. asimpl. move => *. do 2 eexists. split; eauto.
+      scongruence.
+    - move => a ξ [] //=.
       move => u u0 hξ [].
       case : u => //=.
       case : u0 => //=.
       move => p p0 p1 p2 [? ?] [? h]. subst.
       have ? : p0 = p2 by eauto using ren_injective. subst.
       hauto l:on.
-    - move => n a0 a1 ha iha m ξ []//= p hξ [?]. subst.
-      sauto lq:on use:up_injective.
-    - move => n p A B0 B1 hB ihB m ξ + hξ.
+    - move => a0 a1 ha iha ξ []//= p hξ [?]. subst.
+      fcrush use:up_injective.
+    - move => p A B0 B1 hB ihB ξ + hξ.
       case => //= p' A2 B2 [*]. subst.
-      have : (forall i j, (upRen_PTm_PTm) i = (upRen_PTm_PTm) j -> i = j) by sauto.
+      have : (forall i j, (upRen_PTm_PTm ξ) i = (upRen_PTm_PTm ξ) j -> i = j) by sfirstorder use:up_injective.
       move => {}/ihB => ihB.
       spec_refl.
       sauto lq:on.
-  Admitted.
+    - move => P0 P1 a b c hP ihP ξ []//=.
+      move => > /up_injective.
+      hauto q:on ctrs:R.
+    - move => P a b c0 c1 hc ihc ξ []//=.
+      move => > /up_injective /up_injective.
+      hauto q:on ctrs:R.
+  Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    R a b -> R (subst_PTm a) (subst_PTm b).
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
-    move => h. move : m ρ. elim : n a b /h => n.
-    move => a m ρ /=.
+    move => h. move : ρ. elim : a b /h.
+    move => a ρ /=.
     eapply AppEta'; eauto. by asimpl.
     all : hauto lq:on ctrs:R.
   Qed.
 
-  Lemma nf_preservation n (a b : PTm) :
+  Lemma nf_preservation (a b : PTm) :
     ERed.R a b -> (nf a -> nf b) /\ (ne a -> ne b).
   Proof.
     move => h.
-    elim : n a b /h => //=;
-                        hauto lqb:on use:ne_nf_ren,ne_nf.
+    elim : a b /h => //=;
+      hauto lqb:on use:ne_nf_ren,ne_nf.
   Qed.
 
 End ERed.
@@ -1904,41 +1928,40 @@ Module EReds.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AbsCong n (a b : PTm (S n)) :
+  Lemma AbsCong (a b : PTm) :
     rtc ERed.R a b ->
     rtc ERed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
+  Lemma AppCong (a0 a1 b0 b1 : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
     rtc ERed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
+  Lemma PairCong (a0 a1 b0 b1 : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
     rtc ERed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm) :
+  Lemma ProjCong p (a0 a1  : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
+  Lemma BindCong p (A0 A1 : PTm) B0 B1 :
     rtc ERed.R A0 A1 ->
     rtc ERed.R B0 B1 ->
     rtc ERed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
-
-  Lemma SucCong n (a0 a1 : PTm) :
+  Lemma SucCong (a0 a1 : PTm) :
     rtc ERed.R a0 a1 ->
     rtc ERed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
+  Lemma IndCong P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc ERed.R P0 P1 ->
     rtc ERed.R a0 a1 ->
     rtc ERed.R b0 b1 ->
@@ -1946,37 +1969,37 @@ Module EReds.
     rtc ERed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma renaming n m (a b : PTm) (ξ : fin n -> fin m) :
-    rtc ERed.R a b -> rtc ERed.R (ren_PTm a) (ren_PTm b).
+  Lemma renaming (a b : PTm) (ξ : nat -> nat) :
+    rtc ERed.R a b -> rtc ERed.R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof. induction 1; hauto l:on use:ERed.renaming ctrs:rtc. Qed.
 
-  Lemma FromEPar n (a b : PTm) :
+  Lemma FromEPar (a b : PTm) :
     EPar.R a b ->
     rtc ERed.R a b.
   Proof.
-    move => h. elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong, IndCong, SucCong.
-    - move => n a0 a1 _ h.
+    move => h. elim : a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong, IndCong, SucCong.
+    - move => a0 a1 _ h.
       have {}h : rtc ERed.R (ren_PTm shift a0) (ren_PTm shift a1) by apply renaming.
       apply : rtc_r. apply AbsCong. apply AppCong; eauto. apply rtc_refl.
       apply ERed.AppEta.
-    - move => n a0 a1 _ h.
+    - move => a0 a1 _ h.
       apply : rtc_r.
       apply PairCong; eauto using ProjCong.
       apply ERed.PairEta.
   Qed.
 
-  Lemma FromEPars n (a b : PTm) :
+  Lemma FromEPars (a b : PTm) :
     rtc EPar.R a b ->
     rtc ERed.R a b.
   Proof. induction 1; hauto l:on use:FromEPar, @relations.rtc_transitive. Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    rtc ERed.R a b -> rtc ERed.R (subst_PTm a) (subst_PTm b).
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    rtc ERed.R a b -> rtc ERed.R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     induction 1; hauto lq:on ctrs:rtc use:ERed.substing.
   Qed.
 
-  Lemma app_inv n (a0 b0 C : PTm) :
+  Lemma app_inv (a0 b0 C : PTm) :
     rtc ERed.R (PApp a0 b0) C ->
     exists a1 b1, C = PApp a1 b1 /\
                rtc ERed.R a0 a1 /\
@@ -1986,11 +2009,12 @@ Module EReds.
     move : a0 b0 E.
     elim : u C / hu.
     - hauto lq:on ctrs:rtc.
-    - move => a b c ha ha' iha a0 b0 ?. subst.
+    - move => a0 a1 a2 ha ha0 iha b0 b1 ?. subst.
+      inversion ha; subst; spec_refl;
       hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R.
   Qed.
 
-  Lemma proj_inv n p (a C : PTm) :
+  Lemma proj_inv p (a C : PTm) :
     rtc ERed.R (PProj p a) C ->
     exists c, C = PProj p c /\ rtc ERed.R a c.
   Proof.
@@ -2000,7 +2024,7 @@ Module EReds.
       scrush ctrs:rtc,ERed.R inv:ERed.R.
   Qed.
 
-  Lemma bind_inv n p (A : PTm) B C :
+  Lemma bind_inv p (A : PTm) B C :
     rtc ERed.R (PBind p A B) C ->
     exists A0 B0, C = PBind p A0 B0 /\ rtc ERed.R A A0 /\ rtc ERed.R B B0.
   Proof.
@@ -2011,7 +2035,7 @@ Module EReds.
     hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
-  Lemma suc_inv n (a : PTm) C :
+  Lemma suc_inv (a : PTm) C :
     rtc ERed.R (PSuc a) C ->
     exists b, rtc ERed.R a b /\ C = PSuc b.
   Proof.
@@ -2022,14 +2046,14 @@ Module EReds.
     - hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
-  Lemma zero_inv n (C : PTm) :
+  Lemma zero_inv (C : PTm) :
     rtc ERed.R PZero C -> C = PZero.
     move E : PZero => u hu.
     move : E. elim : u C /hu=>//=.
     - hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
-  Lemma ind_inv n P (a : PTm) b c C :
+  Lemma ind_inv P (a : PTm) b c C :
     rtc ERed.R (PInd P a b c) C ->
     exists P0 a0 b0 c0,  rtc ERed.R P P0 /\ rtc ERed.R a a0 /\
                       rtc ERed.R b b0 /\ rtc ERed.R c c0 /\
@@ -2039,7 +2063,7 @@ Module EReds.
     move : P a b c E.
     elim : u C / hu.
     - hauto lq:on ctrs:rtc.
-    - hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
+    - scrush ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
 End EReds.
@@ -2049,35 +2073,35 @@ End EReds.
 Module EJoin.
   Definition R (a b : PTm) := exists c, rtc ERed.R a c /\ rtc ERed.R b c.
 
-  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm) :
+  Lemma hne_app_inj (a0 b0 a1 b1 : PTm) :
     R (PApp a0 b0) (PApp a1 b1) ->
     R a0 a1 /\ R b0 b1.
   Proof.
     hauto lq:on use:EReds.app_inv.
   Qed.
 
-  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm) :
+  Lemma hne_proj_inj p0 p1 (a0 a1 : PTm) :
     R (PProj p0 a0) (PProj p1 a1) ->
     p0 = p1 /\ R a0 a1.
   Proof.
     hauto lq:on rew:off use:EReds.proj_inv.
   Qed.
 
-  Lemma bind_inj n p0 p1 (A0 A1 : PTm) B0 B1 :
+  Lemma bind_inj p0 p1 (A0 A1 : PTm) B0 B1 :
     R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ R A0 A1 /\ R B0 B1.
   Proof.
     hauto lq:on rew:off use:EReds.bind_inv.
   Qed.
 
-  Lemma suc_inj n  (A0 A1 : PTm)  :
+  Lemma suc_inj (A0 A1 : PTm)  :
     R (PSuc A0) (PSuc A1) ->
     R A0 A1.
   Proof.
     hauto lq:on rew:off use:EReds.suc_inv.
   Qed.
 
-  Lemma ind_inj n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
+  Lemma ind_inj P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
     R P0 P1 /\ R a0 a1 /\ R b0 b1 /\ R c0 c1.
   Proof. hauto q:on use:EReds.ind_inv. Qed.
@@ -2094,7 +2118,7 @@ Module RERed.
     R (PProj p (PPair a b)) (if p is PL then a else b)
 
   | IndZero P b c :
-      R (PInd P PZero b c) b
+    R (PInd P PZero b c) b
 
   | IndSuc P a b c :
     R (PInd P (PSuc a) b c) (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
@@ -2224,8 +2248,8 @@ Module REReds.
   Proof. induction 1; hauto lq:on ctrs:rtc use:RERed.FromEta. Qed.
 
   #[local]Ltac solve_s_rec :=
-  move => *; eapply rtc_l; eauto;
-         hauto lq:on ctrs:RERed.R.
+    move => *; eapply rtc_l; eauto;
+           hauto lq:on ctrs:RERed.R.
 
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
@@ -2324,7 +2348,7 @@ Module REReds.
     move : a0 b0 E.
     elim : u C / hu.
     - hauto lq:on ctrs:rtc.
-    - move => a b c ha ha' iha a0 b0 ?. subst.
+    - move => a b c ha ha iha a0 b0 ?. subst.
       hauto lq:on rew:off ctrs:rtc, RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 

From 7f38544a1e05ca3b8ce486ac07eba559ee219403 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 2 Mar 2025 22:41:15 -0500
Subject: [PATCH 04/24] Finish refactoring fp_red

---
 theories/fp_red.v | 441 +++++++++++++++++++++++-----------------------
 1 file changed, 224 insertions(+), 217 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 1e7863b..6aa7cb6 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -12,6 +12,14 @@ Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 Require Import Btauto.
 Require Import Cdcl.Itauto.
 
+Lemma subst_scons_id (a : PTm) :
+  subst_PTm (scons (VarPTm 0) (funcomp VarPTm shift)) a = a.
+Proof.
+  have E : subst_PTm VarPTm a = a by asimpl.
+  rewrite -{2}E.
+  apply ext_PTm. case => //=.
+Qed.
+
 Ltac2 spec_refl () :=
   List.iter
     (fun a => match a with
@@ -2170,79 +2178,79 @@ Module RERed.
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
 
-  Lemma ToBetaEta n (a b : PTm) :
+  Lemma ToBetaEta (a b : PTm) :
     R a b ->
     ERed.R a b \/ RRed.R a b.
   Proof. induction 1; hauto lq:on db:red. Qed.
 
-  Lemma FromBeta n (a b : PTm) :
+  Lemma FromBeta (a b : PTm) :
     RRed.R a b -> RERed.R a b.
   Proof. induction 1; qauto l:on ctrs:R. Qed.
 
-  Lemma FromEta n (a b : PTm) :
+  Lemma FromEta (a b : PTm) :
     ERed.R a b -> RERed.R a b.
   Proof. induction 1; qauto l:on ctrs:R. Qed.
 
-  Lemma ToBetaEtaPar n (a b : PTm) :
+  Lemma ToBetaEtaPar (a b : PTm) :
     R a b ->
     EPar.R a b \/ RRed.R a b.
   Proof.
     hauto q:on use:ERed.ToEPar, ToBetaEta.
   Qed.
 
-  Lemma sn_preservation n (a b : PTm) :
+  Lemma sn_preservation (a b : PTm) :
     R a b ->
     SN a ->
     SN b.
   Proof. hauto q:on use:ToBetaEtaPar, epar_sn_preservation, red_sn_preservation, RPar.FromRRed. Qed.
 
-  Lemma bind_preservation n (a b : PTm) :
+  Lemma bind_preservation (a b : PTm) :
     R a b -> isbind a -> isbind b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma univ_preservation n (a b : PTm) :
+  Lemma univ_preservation (a b : PTm) :
     R a b -> isuniv a -> isuniv b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma nat_preservation n (a b : PTm) :
+  Lemma nat_preservation (a b : PTm) :
     R a b -> isnat a -> isnat b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma sne_preservation n (a b : PTm) :
+  Lemma sne_preservation (a b : PTm) :
     R a b -> SNe a -> SNe b.
   Proof.
     hauto q:on use:ToBetaEtaPar, RPar.FromRRed use:red_sn_preservation, epar_sn_preservation.
   Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    RERed.R a b -> RERed.R (subst_PTm a) (subst_PTm b).
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    RERed.R a b -> RERed.R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     hauto q:on use:ToBetaEta, FromBeta, FromEta, RRed.substing, ERed.substing.
   Qed.
 
-  Lemma hne_preservation n (a b : PTm) :
+  Lemma hne_preservation (a b : PTm) :
     RERed.R a b -> ishne a -> ishne b.
   Proof.  induction 1; sfirstorder. Qed.
 
 End RERed.
 
 Module REReds.
-  Lemma hne_preservation n (a b : PTm) :
+  Lemma hne_preservation (a b : PTm) :
     rtc RERed.R a b -> ishne a -> ishne b.
   Proof.  induction 1; eauto using RERed.hne_preservation, rtc_refl, rtc. Qed.
 
-  Lemma sn_preservation n (a b : PTm) :
+  Lemma sn_preservation (a b : PTm) :
     rtc RERed.R a b ->
     SN a ->
     SN b.
   Proof. induction 1; eauto using RERed.sn_preservation. Qed.
 
-  Lemma FromRReds n (a b : PTm) :
+  Lemma FromRReds (a b : PTm) :
     rtc RRed.R a b ->
     rtc RERed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:RERed.FromBeta. Qed.
 
-  Lemma FromEReds n (a b : PTm) :
+  Lemma FromEReds (a b : PTm) :
     rtc ERed.R a b ->
     rtc RERed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:RERed.FromEta. Qed.
@@ -2254,34 +2262,34 @@ Module REReds.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); apply rtc_refl.
 
-  Lemma AbsCong n (a b : PTm (S n)) :
+  Lemma AbsCong (a b : PTm) :
     rtc RERed.R a b ->
     rtc RERed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
+  Lemma AppCong (a0 a1 b0 b1 : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R b0 b1 ->
     rtc RERed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
+  Lemma PairCong (a0 a1 b0 b1 : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R b0 b1 ->
     rtc RERed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm) :
+  Lemma ProjCong p (a0 a1  : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm) :
+  Lemma SucCong (a0 a1 : PTm) :
     rtc RERed.R a0 a1 ->
     rtc RERed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
+  Lemma IndCong P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc RERed.R P0 P1 ->
     rtc RERed.R a0 a1 ->
     rtc RERed.R b0 b1 ->
@@ -2289,30 +2297,30 @@ Module REReds.
     rtc RERed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
+  Lemma BindCong p (A0 A1 : PTm) B0 B1 :
     rtc RERed.R A0 A1 ->
     rtc RERed.R B0 B1 ->
     rtc RERed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
-  Lemma bind_preservation n (a b : PTm) :
+  Lemma bind_preservation (a b : PTm) :
     rtc RERed.R a b -> isbind a -> isbind b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.bind_preservation. Qed.
 
-  Lemma univ_preservation n (a b : PTm) :
+  Lemma univ_preservation (a b : PTm) :
     rtc RERed.R a b -> isuniv a -> isuniv b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.univ_preservation. Qed.
 
-  Lemma nat_preservation n (a b : PTm) :
+  Lemma nat_preservation (a b : PTm) :
     rtc RERed.R a b -> isnat a -> isnat b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.nat_preservation. Qed.
 
-  Lemma sne_preservation n (a b : PTm) :
+  Lemma sne_preservation (a b : PTm) :
     rtc RERed.R a b -> SNe a -> SNe b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.sne_preservation. Qed.
 
-  Lemma bind_inv n p A B C :
-    rtc (@RERed.R n) (PBind p A B) C ->
+  Lemma bind_inv p A B C :
+    rtc RERed.R(PBind p A B) C ->
     exists A0 B0, C = PBind p A0 B0 /\ rtc RERed.R A A0 /\ rtc RERed.R B B0.
   Proof.
     move E : (PBind p A B) => u hu.
@@ -2320,8 +2328,8 @@ Module REReds.
     elim : u C / hu; sauto lq:on rew:off.
   Qed.
 
-  Lemma univ_inv n i C :
-    rtc (@RERed.R n) (PUniv i) C  ->
+  Lemma univ_inv i C :
+    rtc RERed.R (PUniv i) C  ->
     C = PUniv i.
   Proof.
     move E : (PUniv i) => u hu.
@@ -2329,7 +2337,7 @@ Module REReds.
     hauto lq:on rew:off ctrs:rtc inv:RERed.R.
   Qed.
 
-  Lemma var_inv n (i : fin n) C :
+  Lemma var_inv (i : nat) C :
     rtc RERed.R (VarPTm i) C ->
     C = VarPTm i.
   Proof.
@@ -2337,7 +2345,7 @@ Module REReds.
     move : i E. elim : u C /hu; hauto lq:on rew:off inv:RERed.R.
   Qed.
 
-  Lemma hne_app_inv n (a0 b0 C : PTm) :
+  Lemma hne_app_inv (a0 b0 C : PTm) :
     rtc RERed.R (PApp a0 b0) C ->
     ishne a0 ->
     exists a1 b1, C = PApp a1 b1 /\
@@ -2348,11 +2356,11 @@ Module REReds.
     move : a0 b0 E.
     elim : u C / hu.
     - hauto lq:on ctrs:rtc.
-    - move => a b c ha ha iha a0 b0 ?. subst.
+    - move => a b c ha0 ha1 iha a0 b0 ?. subst.
       hauto lq:on rew:off ctrs:rtc, RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 
-  Lemma hne_proj_inv n p (a C : PTm) :
+  Lemma hne_proj_inv p (a C : PTm) :
     rtc RERed.R (PProj p a) C ->
     ishne a ->
     exists c, C = PProj p c /\ rtc RERed.R a c.
@@ -2363,7 +2371,7 @@ Module REReds.
       scrush ctrs:rtc,RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 
-  Lemma hne_ind_inv n P a b c (C : PTm) :
+  Lemma hne_ind_inv P a b c (C : PTm) :
     rtc RERed.R (PInd P a b c) C -> ishne a ->
     exists P0 a0 b0 c0, C = PInd P0 a0 b0 c0 /\
                      rtc RERed.R P P0 /\
@@ -2377,47 +2385,47 @@ Module REReds.
       scrush ctrs:rtc,RERed.R use:RERed.hne_preservation inv:RERed.R.
   Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    rtc RERed.R a b -> rtc RERed.R (subst_PTm a) (subst_PTm b).
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    rtc RERed.R a b -> rtc RERed.R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     induction 1; hauto lq:on ctrs:rtc use:RERed.substing.
   Qed.
 
 
-  Lemma cong_up n m (ρ0 ρ1 : fin n -> PTm) :
+  Lemma cong_up (ρ0 ρ1 : nat -> PTm) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    (forall i, rtc RERed.R (up_PTm_PTm0 i) (up_PTm_PTm1 i)).
-  Proof. move => h i. destruct i as [i|].
+    (forall i, rtc RERed.R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)).
+  Proof. move => h [|i]; cycle 1.
          simpl. rewrite /funcomp.
          substify. by apply substing.
          apply rtc_refl.
   Qed.
 
-  Lemma cong_up2 n m (ρ0 ρ1 : fin n -> PTm) :
+  Lemma cong_up2 (ρ0 ρ1 : nat -> PTm) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    (forall i, rtc RERed.R (up_PTm_PTm (up_PTm_PTm0) i) (up_PTm_PTm (up_PTm_PTm1) i)).
+    (forall i, rtc RERed.R (up_PTm_PTm (up_PTm_PTm ρ0) i) (up_PTm_PTm (up_PTm_PTm ρ1) i)).
   Proof. hauto l:on use:cong_up. Qed.
 
-  Lemma cong n m (a : PTm) (ρ0 ρ1 : fin n -> PTm) :
+  Lemma cong (a : PTm) (ρ0 ρ1 : nat -> PTm) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    rtc RERed.R (subst_PTm0 a) (subst_PTm1 a).
+    rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 a).
   Proof.
-    move : m ρ0 ρ1. elim : n / a => /=;
+    move : ρ0 ρ1. elim : a => /=;
       eauto 20 using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl, IndCong, SucCong, cong_up2.
   Qed.
 
-  Lemma cong' n m (a b : PTm) (ρ0 ρ1 : fin n -> PTm) :
+  Lemma cong' (a b : PTm) (ρ0 ρ1 : nat -> PTm) :
     rtc RERed.R a b ->
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
-    rtc RERed.R (subst_PTm0 a) (subst_PTm1 b).
+    rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 b).
   Proof.
     move => h0 h1.
-    have : rtc RERed.R (subst_PTm0 a) (subst_PTm1 a) by eauto using cong.
+    have : rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 a) by eauto using cong.
     move => ?. apply : relations.rtc_transitive; eauto.
     hauto l:on use:substing.
   Qed.
 
-  Lemma ToEReds n (a b : PTm) :
+  Lemma ToEReds (a b : PTm) :
     nf a ->
     rtc RERed.R a b -> rtc ERed.R a b.
   Proof.
@@ -2425,14 +2433,14 @@ Module REReds.
     induction h; hauto lq:on rew:off ctrs:rtc use:RERed.ToBetaEta, RReds.nf_refl, @rtc_once, ERed.nf_preservation.
   Qed.
 
-  Lemma zero_inv n (C : PTm) :
+  Lemma zero_inv (C : PTm) :
     rtc RERed.R PZero C -> C = PZero.
     move E : PZero => u hu.
     move : E. elim : u C /hu=>//=.
     - hauto lq:on rew:off ctrs:rtc, RERed.R inv:RERed.R, rtc.
   Qed.
 
-  Lemma suc_inv n (a : PTm) C :
+  Lemma suc_inv (a : PTm) C :
     rtc RERed.R (PSuc a) C ->
     exists b, rtc RERed.R a b /\ C = PSuc b.
   Proof.
@@ -2443,8 +2451,8 @@ Module REReds.
     - hauto lq:on rew:off ctrs:rtc, RERed.R inv:RERed.R, rtc.
   Qed.
 
-  Lemma nat_inv n C :
-    rtc RERed.R (@PNat n) C ->
+  Lemma nat_inv C :
+    rtc RERed.R PNat C ->
     C = PNat.
   Proof.
     move E : PNat => u hu. move : E.
@@ -2523,36 +2531,36 @@ Module LoRed.
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
 
-  Lemma hf_preservation n (a b : PTm) :
+  Lemma hf_preservation (a b : PTm) :
     LoRed.R a b ->
     ishf a ->
     ishf b.
   Proof.
-    move => h. elim : n a b /h=>//=.
+    move => h. elim : a b /h=>//=.
   Qed.
 
-  Lemma ToRRed n (a b : PTm) :
+  Lemma ToRRed (a b : PTm) :
     LoRed.R a b ->
     RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
 
-  Lemma AppAbs' n a (b : PTm) u :
+  Lemma AppAbs' a (b : PTm) u :
     u = (subst_PTm (scons b VarPTm) a) ->
     R (PApp (PAbs a) b)  u.
   Proof. move => ->. apply AppAbs. Qed.
 
-  Lemma IndSuc' n P a b c u :
+  Lemma IndSuc' P a b c u :
     u = (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ->
-    R (@PInd n P (PSuc a) b c) u.
+    R (@PInd P (PSuc a) b c) u.
   Proof. move => ->. apply IndSuc. Qed.
 
-  Lemma renaming n m (a b : PTm) (ξ : fin n -> fin m) :
-    R a b -> R (ren_PTm a) (ren_PTm b).
+  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 => n a b m ξ /=.
+    move => a b ξ /=.
     apply AppAbs'. by asimpl.
     all : try qauto ctrs:R use:ne_nf_ren, ishf_ren.
     move => * /=; apply IndSuc'. by asimpl.
@@ -2561,7 +2569,7 @@ Module LoRed.
 End LoRed.
 
 Module LoReds.
-  Lemma hf_preservation n (a b : PTm) :
+  Lemma hf_preservation (a b : PTm) :
     rtc LoRed.R a b ->
     ishf a ->
     ishf b.
@@ -2569,7 +2577,7 @@ Module LoReds.
     induction 1; eauto using LoRed.hf_preservation.
   Qed.
 
-  Lemma hf_ne_imp n (a b : PTm) :
+  Lemma hf_ne_imp (a b : PTm) :
     rtc LoRed.R a b ->
     ne b ->
     ~~ ishf a.
@@ -2585,39 +2593,39 @@ Module LoReds.
   #[local]Ltac solve_s :=
     repeat (induction 1; last by solve_s_rec); (move => *; apply rtc_refl).
 
-  Lemma AbsCong n (a b : PTm (S n)) :
+  Lemma AbsCong (a b : PTm) :
     rtc LoRed.R a b ->
     rtc LoRed.R (PAbs a) (PAbs b).
   Proof. solve_s. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
+  Lemma AppCong (a0 a1 b0 b1 : PTm) :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R b0 b1 ->
     ne a1 ->
     rtc LoRed.R (PApp a0 b0) (PApp a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
+  Lemma PairCong (a0 a1 b0 b1 : PTm) :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R b0 b1 ->
     nf a1 ->
     rtc LoRed.R (PPair a0 b0) (PPair a1 b1).
   Proof. solve_s. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm) :
+  Lemma ProjCong p (a0 a1  : PTm) :
     rtc LoRed.R a0 a1 ->
     ne a1 ->
     rtc LoRed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
+  Lemma BindCong p (A0 A1 : PTm) B0 B1 :
     rtc LoRed.R A0 A1 ->
     rtc LoRed.R B0 B1 ->
     nf A1 ->
     rtc LoRed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
+  Lemma IndCong P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R P0 P1 ->
     rtc LoRed.R b0 b1 ->
@@ -2628,14 +2636,14 @@ Module LoReds.
     rtc LoRed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. solve_s. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm) :
+  Lemma SucCong (a0 a1 : PTm) :
     rtc LoRed.R a0 a1 ->
     rtc LoRed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
   Local Ltac triv := simpl in *; itauto.
 
-  Lemma FromSN_mutual : forall n,
+  Lemma FromSN_mutual :
     (forall (a : PTm) (_ : SNe a), exists v, rtc LoRed.R a v /\ ne v) /\
     (forall (a : PTm) (_ : SN a), exists v, rtc LoRed.R a v /\ nf v) /\
     (forall (a b : PTm) (_ : TRedSN a b), LoRed.R a b).
@@ -2644,7 +2652,6 @@ Module LoReds.
     - hauto lq:on ctrs:rtc.
     - hauto lq:on rew:off use:LoReds.AppCong solve+:triv.
     - hauto l:on use:LoReds.ProjCong solve+:triv.
-    - hauto lq:on ctrs:rtc.
     - hauto lq:on use:LoReds.IndCong solve+:triv.
     - hauto q:on use:LoReds.PairCong solve+:triv.
     - hauto q:on use:LoReds.AbsCong solve+:triv.
@@ -2656,73 +2663,71 @@ Module LoReds.
     - hauto lq:on ctrs:rtc.
     - hauto l:on use:SucCong.
     - qauto ctrs:LoRed.R.
-    - move => n a0 a1 b hb ihb h.
+    - move => a0 a1 b hb ihb h.
       have : ~~ ishf a0 by inversion h.
       hauto lq:on ctrs:LoRed.R.
     - qauto ctrs:LoRed.R.
     - qauto ctrs:LoRed.R.
-    - move => n p a b h.
+    - move => p a b h.
       have : ~~ ishf a by inversion h.
       hauto lq:on ctrs:LoRed.R.
     - sfirstorder.
     - sfirstorder.
-    - move => n P a0 a1 b c hP ihP hb ihb hc ihc hr.
+    - move => P a0 a1 b c hP ihP hb ihb hc ihc hr.
       have : ~~ ishf a0 by inversion hr.
       hauto q:on ctrs:LoRed.R.
   Qed.
 
-  Lemma FromSN : forall n a, @SN n a -> exists v, rtc LoRed.R a v /\ nf v.
+  Lemma FromSN : forall a, SN a -> exists v, rtc LoRed.R a v /\ nf v.
   Proof. firstorder using FromSN_mutual. Qed.
 
-  Lemma ToRReds : forall n (a b : PTm), rtc LoRed.R a b -> rtc RRed.R a b.
+  Lemma ToRReds : forall (a b : PTm), rtc LoRed.R a b -> rtc RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:rtc use:LoRed.ToRRed. Qed.
-
 End LoReds.
 
 
 Fixpoint size_PTm (a : PTm) :=
   match a with
   | VarPTm _ => 1
-  | PAbs a => 3 + size_PTm
-  | PApp a b => 1 + Nat.add (size_PTm) (size_PTm)
-  | PProj p a => 1 + size_PTm
-  | PPair a b => 3 + Nat.add (size_PTm) (size_PTm)
+  | PAbs a => 3 + size_PTm a
+  | PApp a b => 1 + Nat.add (size_PTm a) (size_PTm b)
+  | PProj p a => 1 + size_PTm a
+  | PPair a b => 3 + Nat.add (size_PTm a) (size_PTm b)
   | PUniv _ => 3
-  | PBind p A B => 3 + Nat.add (size_PTm) (size_PTm)
-  | PInd P a b c => 3 + size_PTm + size_PTm + size_PTm + size_PTm
+  | PBind p A B => 3 + Nat.add (size_PTm A) (size_PTm B)
+  | PInd P a b c => 3 + size_PTm P + size_PTm a + size_PTm b + size_PTm c
   | PNat => 3
-  | PSuc a => 3 + size_PTm
+  | PSuc a => 3 + size_PTm a
   | PZero => 3
-  | PBot => 1
   end.
 
-Lemma size_PTm_ren n m (ξ : fin n -> fin m) a : size_PTm (ren_PTm a) = size_PTm.
+Lemma size_PTm_ren (ξ : nat -> nat) a : size_PTm (ren_PTm ξ a) = size_PTm a.
 Proof.
-  move : m ξ. elim : n / a => //=; scongruence.
+  move : ξ. elim : a => //=; scongruence.
 Qed.
 
 #[export]Hint Rewrite size_PTm_ren : sizetm.
 
 Lemma ered_size (a b : PTm) :
   ERed.R a b ->
-  size_PTm < size_PTm.
+  size_PTm b < size_PTm a.
 Proof.
-  move => h. elim : n a b /h; hauto l:on rew:db:sizetm.
+  move => h. elim : a b /h; try hauto l:on rew:db:sizetm solve+:lia.
 Qed.
 
-Lemma ered_sn n (a : PTm) : sn ERed.R a.
+Lemma ered_sn (a : PTm) : sn ERed.R a.
 Proof.
   hauto lq:on rew:off use:size_PTm_ren, ered_size,
           well_founded_lt_compat unfold:well_founded.
 Qed.
 
-Lemma ered_local_confluence n (a b c : PTm) :
+Lemma ered_local_confluence (a b c : PTm) :
   ERed.R a b ->
   ERed.R a c ->
   exists d, rtc ERed.R b d  /\ rtc ERed.R c d.
 Proof.
   move => h. move : c.
-  elim : n a b / h => n.
+  elim : a b / h.
   - move => a c.
     elim /ERed.inv => //= _.
     + move => a0 [+ ?]. subst => h.
@@ -2822,7 +2827,7 @@ Proof.
   - qauto l:on inv:ERed.R ctrs:rtc use:EReds.SucCong.
 Qed.
 
-Lemma ered_confluence n (a b c : PTm) :
+Lemma ered_confluence (a b c : PTm) :
   rtc ERed.R a b ->
   rtc ERed.R a c ->
   exists d, rtc ERed.R b d  /\ rtc ERed.R c d.
@@ -2830,7 +2835,7 @@ Proof.
   sfirstorder use:relations.locally_confluent_confluent, ered_sn, ered_local_confluence.
 Qed.
 
-Lemma red_confluence n (a b c : PTm) :
+Lemma red_confluence (a b c : PTm) :
   rtc RRed.R a b ->
   rtc RRed.R a c ->
   exists d, rtc RRed.R b d  /\ rtc RRed.R c d.
@@ -2841,7 +2846,7 @@ Lemma red_confluence n (a b c : PTm) :
   eauto using RPar.diamond.
 Qed.
 
-Lemma red_uniquenf n (a b c : PTm) :
+Lemma red_uniquenf (a b c : PTm) :
   rtc RRed.R a b ->
   rtc RRed.R a c ->
   nf b ->
@@ -2856,20 +2861,20 @@ Proof.
 Qed.
 
 Module NeEPars.
-  Lemma R_nonelim_nf n (a b : PTm) :
+  Lemma R_nonelim_nf (a b : PTm) :
     rtc NeEPar.R_nonelim a b ->
     nf b ->
     nf a.
   Proof. induction 1; sfirstorder use:NeEPar.R_elim_nf. Qed.
 
-  Lemma ToEReds : forall n, (forall (a b : PTm), rtc NeEPar.R_nonelim a b -> rtc ERed.R a b).
+  Lemma ToEReds : (forall (a b : PTm), rtc NeEPar.R_nonelim a b -> rtc ERed.R a b).
   Proof.
     induction 1; hauto l:on use:NeEPar.ToEPar, EReds.FromEPar, @relations.rtc_transitive.
   Qed.
 End NeEPars.
 
 
-Lemma rered_standardization n (a c : PTm) :
+Lemma rered_standardization (a c : PTm) :
   SN a ->
   rtc RERed.R a c ->
   exists b, rtc RRed.R a b /\ rtc NeEPar.R_nonelim b c.
@@ -2886,7 +2891,7 @@ Proof.
   - hauto lq:on ctrs:rtc use:red_sn_preservation, RPar.FromRRed.
 Qed.
 
-Lemma rered_confluence n (a b c : PTm) :
+Lemma rered_confluence (a b c : PTm) :
   SN a ->
   rtc RERed.R a b ->
   rtc RERed.R a c ->
@@ -2921,13 +2926,13 @@ Qed.
 (* Beta joinability *)
 Module BJoin.
   Definition R (a b : PTm) := exists c, rtc RRed.R a c /\ rtc RRed.R b c.
-  Lemma refl n (a : PTm) : R a a.
+  Lemma refl (a : PTm) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
-  Lemma symmetric n (a b : PTm) : R a b -> R b a.
+  Lemma symmetric (a b : PTm) : R a b -> R b a.
   Proof. sfirstorder unfold:R. Qed.
 
-  Lemma transitive n (a b c : PTm) : R a b -> R b c -> R a c.
+  Lemma transitive (a b c : PTm) : R a b -> R b c -> R a c.
   Proof.
     rewrite /R.
     move => [ab [ha +]] [bc [+ hc]].
@@ -2951,19 +2956,19 @@ End BJoin.
 Module DJoin.
   Definition R (a b : PTm) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.
 
-  Lemma refl n (a : PTm) : R a a.
+  Lemma refl (a : PTm) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
-  Lemma FromEJoin n (a b : PTm) : EJoin.R a b -> DJoin.R a b.
+  Lemma FromEJoin (a b : PTm) : EJoin.R a b -> DJoin.R a b.
   Proof. hauto q:on use:REReds.FromEReds. Qed.
 
-  Lemma ToEJoin n (a b : PTm) : nf a -> nf b -> DJoin.R a b -> EJoin.R a b.
+  Lemma ToEJoin (a b : PTm) : nf a -> nf b -> DJoin.R a b -> EJoin.R a b.
   Proof. hauto q:on use:REReds.ToEReds. Qed.
 
-  Lemma symmetric n (a b : PTm) : R a b -> R b a.
+  Lemma symmetric (a b : PTm) : R a b -> R b a.
   Proof. sfirstorder unfold:R. Qed.
 
-  Lemma transitive n (a b c : PTm) : SN b -> R a b -> R b c -> R a c.
+  Lemma transitive (a b c : PTm) : SN b -> R a b -> R b c -> R a c.
   Proof.
     rewrite /R.
     move => + [ab [ha +]] [bc [+ hc]].
@@ -2972,35 +2977,35 @@ Module DJoin.
     exists v. sfirstorder use:@relations.rtc_transitive.
   Qed.
 
-  Lemma AbsCong n (a b : PTm (S n)) :
+  Lemma AbsCong (a b : PTm) :
     R a b ->
     R (PAbs a) (PAbs b).
   Proof. hauto lq:on use:REReds.AbsCong unfold:R. Qed.
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm) :
+  Lemma AppCong (a0 a1 b0 b1 : PTm) :
     R a0 a1 ->
     R b0 b1 ->
     R (PApp a0 b0) (PApp a1 b1).
   Proof. hauto lq:on use:REReds.AppCong unfold:R. Qed.
 
-  Lemma PairCong n (a0 a1 b0 b1 : PTm) :
+  Lemma PairCong (a0 a1 b0 b1 : PTm) :
     R a0 a1 ->
     R b0 b1 ->
     R (PPair a0 b0) (PPair a1 b1).
   Proof. hauto q:on use:REReds.PairCong. Qed.
 
-  Lemma ProjCong n p (a0 a1  : PTm) :
+  Lemma ProjCong p (a0 a1  : PTm) :
     R a0 a1 ->
     R (PProj p a0) (PProj p a1).
   Proof. hauto q:on use:REReds.ProjCong. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm) B0 B1 :
+  Lemma BindCong p (A0 A1 : PTm) B0 B1 :
     R A0 A1 ->
     R B0 B1 ->
     R (PBind p A0 B0) (PBind p A1 B1).
   Proof. hauto q:on use:REReds.BindCong. Qed.
 
-  Lemma IndCong n P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
+  Lemma IndCong P0 P1 (a0 a1 : PTm) b0 b1 c0 c1 :
     R P0 P1 ->
     R a0 a1 ->
     R b0 b1 ->
@@ -3008,12 +3013,12 @@ Module DJoin.
     R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
   Proof. hauto q:on use:REReds.IndCong. Qed.
 
-  Lemma SucCong n (a0 a1 : PTm) :
+  Lemma SucCong (a0 a1 : PTm) :
     R a0 a1 ->
     R (PSuc a0) (PSuc a1).
   Proof. qauto l:on use:REReds.SucCong. Qed.
 
-  Lemma FromRedSNs n (a b : PTm) :
+  Lemma FromRedSNs (a b : PTm) :
     rtc TRedSN a b ->
     R a b.
   Proof.
@@ -3021,7 +3026,7 @@ Module DJoin.
     sfirstorder use:@rtc_refl unfold:R.
   Qed.
 
-  Lemma sne_nat_noconf n (a b : PTm) :
+  Lemma sne_nat_noconf (a b : PTm) :
     R a b -> SNe a -> isnat b -> False.
   Proof.
     move => [c [? ?]] *.
@@ -3029,7 +3034,7 @@ Module DJoin.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma sne_bind_noconf n (a b : PTm) :
+  Lemma sne_bind_noconf (a b : PTm) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
     move => [c [? ?]] *.
@@ -3037,14 +3042,14 @@ Module DJoin.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma sne_univ_noconf n (a b : PTm) :
+  Lemma sne_univ_noconf (a b : PTm) :
     R a b -> SNe a -> isuniv b -> False.
   Proof.
     hauto q:on use:REReds.sne_preservation,
           REReds.univ_preservation inv:SNe.
   Qed.
 
-  Lemma bind_univ_noconf n (a b : PTm) :
+  Lemma bind_univ_noconf (a b : PTm) :
     R a b -> isbind a -> isuniv b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3054,7 +3059,7 @@ Module DJoin.
     case : c => //=; itauto.
   Qed.
 
-  Lemma hne_univ_noconf n (a b : PTm) :
+  Lemma hne_univ_noconf (a b : PTm) :
     R a b -> ishne a -> isuniv b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3065,7 +3070,7 @@ Module DJoin.
     case : c => //=.
   Qed.
 
-  Lemma hne_bind_noconf n (a b : PTm) :
+  Lemma hne_bind_noconf (a b : PTm) :
     R a b -> ishne a -> isbind b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3076,7 +3081,7 @@ Module DJoin.
     case : c => //=.
   Qed.
 
-  Lemma hne_nat_noconf n (a b : PTm) :
+  Lemma hne_nat_noconf (a b : PTm) :
     R a b -> ishne a -> isnat b -> False.
   Proof.
     move => [c [h0 h1]] h2 h3.
@@ -3084,7 +3089,7 @@ Module DJoin.
     clear. case : c => //=; itauto.
   Qed.
 
-  Lemma bind_inj n p0 p1 (A0 A1 : PTm) B0 B1 :
+  Lemma bind_inj p0 p1 (A0 A1 : PTm) B0 B1 :
     DJoin.R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ DJoin.R A0 A1 /\ DJoin.R B0 B1.
   Proof.
@@ -3092,24 +3097,24 @@ Module DJoin.
     hauto lq:on rew:off use:REReds.bind_inv.
   Qed.
 
-  Lemma var_inj n (i j : fin n) :
+  Lemma var_inj (i j : nat) :
     R (VarPTm i) (VarPTm j) -> i = j.
   Proof. sauto lq:on rew:off use:REReds.var_inv unfold:R. Qed.
 
-  Lemma univ_inj n i j :
-    @R n (PUniv i) (PUniv j)  -> i = j.
+  Lemma univ_inj i j :
+    @R (PUniv i) (PUniv j)  -> i = j.
   Proof.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
-  Lemma suc_inj n  (A0 A1 : PTm)  :
+  Lemma suc_inj  (A0 A1 : PTm)  :
     R (PSuc A0) (PSuc A1) ->
     R A0 A1.
   Proof.
     hauto lq:on rew:off use:REReds.suc_inv.
   Qed.
 
-  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm) :
+  Lemma hne_app_inj (a0 b0 a1 b1 : PTm) :
     R (PApp a0 b0) (PApp a1 b1) ->
     ishne a0 ->
     ishne a1 ->
@@ -3118,7 +3123,7 @@ Module DJoin.
     hauto lq:on use:REReds.hne_app_inv.
   Qed.
 
-  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm) :
+  Lemma hne_proj_inj p0 p1 (a0 a1 : PTm) :
     R (PProj p0 a0) (PProj p1 a1) ->
     ishne a0 ->
     ishne a1 ->
@@ -3127,62 +3132,62 @@ Module DJoin.
     sauto lq:on use:REReds.hne_proj_inv.
   Qed.
 
-  Lemma FromRRed0 n (a b : PTm) :
+  Lemma FromRRed0 (a b : PTm) :
     RRed.R a b -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
   Qed.
 
-  Lemma FromRRed1 n (a b : PTm) :
+  Lemma FromRRed1 (a b : PTm) :
     RRed.R b a -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
   Qed.
 
-  Lemma FromRReds n (a b : PTm) :
+  Lemma FromRReds (a b : PTm) :
     rtc RRed.R b a -> R a b.
   Proof.
     hauto lq:on ctrs:rtc use:REReds.FromRReds unfold:R.
   Qed.
 
-  Lemma FromBJoin n (a b : PTm) :
+  Lemma FromBJoin (a b : PTm) :
     BJoin.R a b -> R a b.
   Proof.
-    hauto lq:on ctrs:rtc use:REReds.FromRReds unfold:R.
+    hauto lq:on ctrs:rtc use:REReds.FromRReds unfold:R, BJoin.R.
   Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    R a b -> R (subst_PTm a) (subst_PTm b).
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     hauto lq:on rew:off ctrs:rtc unfold:R use:REReds.substing.
   Qed.
 
-  Lemma renaming n m (a b : PTm) (ρ : fin n -> fin m) :
-    R a b -> R (ren_PTm a) (ren_PTm b).
+  Lemma renaming (a b : PTm) (ρ : nat -> nat) :
+    R a b -> R (ren_PTm ρ a) (ren_PTm ρ b).
   Proof. substify. apply substing. Qed.
 
-  Lemma weakening n (a b : PTm)  :
+  Lemma weakening (a b : PTm)  :
     R a b -> R (ren_PTm shift a) (ren_PTm shift b).
   Proof. apply renaming. Qed.
 
-  Lemma cong n m (a : PTm (S n)) c d (ρ : fin n -> PTm) :
+  Lemma cong (a : PTm) c d (ρ : nat -> PTm) :
     R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) a).
   Proof.
     rewrite /R. move => [cd [h0 h1]].
     exists (subst_PTm (scons cd ρ) a).
-    hauto q:on ctrs:rtc inv:option use:REReds.cong.
+    hauto q:on ctrs:rtc inv:nat use:REReds.cong.
   Qed.
 
-  Lemma cong' n m (a b : PTm (S n)) c d (ρ : fin n -> PTm) :
+  Lemma cong' (a b : PTm) c d (ρ : nat -> PTm) :
     R a b ->
     R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) b).
   Proof.
     rewrite /R. move => [ab [h2 h3]] [cd [h0 h1]].
     exists (subst_PTm (scons cd ρ) ab).
-    hauto q:on ctrs:rtc inv:option use:REReds.cong'.
+    hauto q:on ctrs:rtc inv:nat use:REReds.cong'.
   Qed.
 
-  Lemma pair_inj n (a0 a1 b0 b1 : PTm) :
+  Lemma pair_inj (a0 a1 b0 b1 : PTm) :
     SN (PPair a0 b0) ->
     SN (PPair a1 b1) ->
     R (PPair a0 b0) (PPair a1 b1) ->
@@ -3209,7 +3214,7 @@ Module DJoin.
     split; eauto using transitive.
   Qed.
 
-  Lemma ejoin_pair_inj n (a0 a1 b0 b1 : PTm) :
+  Lemma ejoin_pair_inj (a0 a1 b0 b1 : PTm) :
     nf a0 -> nf b0 -> nf a1 -> nf b1 ->
     EJoin.R (PPair a0 b0) (PPair a1 b1) ->
     EJoin.R a0 a1 /\ EJoin.R b0 b1.
@@ -3221,7 +3226,7 @@ Module DJoin.
     hauto l:on use:ToEJoin.
   Qed.
 
-  Lemma abs_inj n (a0 a1 : PTm (S n)) :
+  Lemma abs_inj (a0 a1 : PTm) :
     SN a0 -> SN a1 ->
     R (PAbs a0) (PAbs a1) ->
     R a0 a1.
@@ -3231,22 +3236,23 @@ Module DJoin.
     move /AppCong. move /(_ (VarPTm var_zero) (VarPTm var_zero)).
     move => /(_ ltac:(sfirstorder use:refl)).
     move => h.
-    have /FromRRed1 h0 : RRed.R (PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a0)) (VarPTm var_zero)) a0 by apply RRed.AppAbs'; asimpl.
-    have /FromRRed0 h1 : RRed.R (PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a1)) (VarPTm var_zero)) a1 by apply RRed.AppAbs'; asimpl.
+    have /FromRRed1 h0 : RRed.R (PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a0)) (VarPTm var_zero)) a0. apply RRed.AppAbs'; asimpl. by rewrite subst_scons_id.
+    have /FromRRed0 h1 : RRed.R (PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a1)) (VarPTm var_zero)) a1 by
+      apply RRed.AppAbs'; eauto; by asimpl; rewrite ?subst_scons_id.
     have sn0' := sn0.
     have sn1' := sn1.
     eapply sn_renaming with (ξ := (upRen_PTm_PTm shift)) in sn0.
     eapply sn_renaming with (ξ := (upRen_PTm_PTm shift)) in sn1.
     apply : transitive; try apply h0.
     apply : N_Exp. apply N_β. sauto.
-    by asimpl.
+    asimpl. by rewrite subst_scons_id.
     apply : transitive; try apply h1.
     apply : N_Exp. apply N_β. sauto.
-    by asimpl.
+    by asimpl; rewrite subst_scons_id.
     eauto.
   Qed.
 
-  Lemma ejoin_abs_inj n (a0 a1 : PTm (S n)) :
+  Lemma ejoin_abs_inj (a0 a1 : PTm) :
     nf a0 -> nf a1 ->
     EJoin.R (PAbs a0) (PAbs a1) ->
     EJoin.R a0 a1.
@@ -3258,13 +3264,13 @@ Module DJoin.
     hauto l:on use:ToEJoin.
   Qed.
 
-  Lemma standardization n (a b : PTm) :
+  Lemma standardization (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc RRed.R a va /\ rtc RRed.R b vb /\ nf va /\ nf vb /\ EJoin.R va vb.
   Proof.
     move => h0 h1 [ab [hv0 hv1]].
     have hv : SN ab by sfirstorder use:REReds.sn_preservation.
-    have : exists v, rtc RRed.R ab v  /\ nf v by hauto q:on use:LoReds.FromSN, LoReds.ToRReds.
+    have : exists v, rtc RRed.R ab v  /\ nf v by sfirstorder use:LoReds.FromSN, LoReds.ToRReds.
     move => [v [hv2 hv3]].
     have : rtc RERed.R a v by hauto q:on use:@relations.rtc_transitive, REReds.FromRReds.
     have : rtc RERed.R b v by hauto q:on use:@relations.rtc_transitive, REReds.FromRReds.
@@ -3278,7 +3284,7 @@ Module DJoin.
     hauto q:on use:NeEPars.ToEReds unfold:EJoin.R.
   Qed.
 
-  Lemma standardization_lo n (a b : PTm) :
+  Lemma standardization_lo (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc LoRed.R a va /\ rtc LoRed.R b vb /\ nf va /\ nf vb /\ EJoin.R va vb.
   Proof.
@@ -3311,21 +3317,21 @@ Module Sub1.
     R B0 B1 ->
     R (PBind PSig A0 B0) (PBind PSig A1 B1).
 
-  Lemma transitive0 n (A B C : PTm) :
+  Lemma transitive0 (A B C : PTm) :
     R A B -> (R B C -> R A C) /\ (R C A -> R C B).
   Proof.
     move => h. move : C.
-    elim : n A B /h; hauto lq:on ctrs:R inv:R solve+:lia.
+    elim : A B /h; hauto lq:on ctrs:R inv:R solve+:lia.
   Qed.
 
-  Lemma transitive n (A B C : PTm) :
+  Lemma transitive (A B C : PTm) :
     R A B -> R B C -> R A C.
   Proof. hauto q:on use:transitive0. Qed.
 
-  Lemma refl n (A : PTm) : R A A.
+  Lemma refl (A : PTm) : R A A.
   Proof. sfirstorder. Qed.
 
-  Lemma commutativity0 n (A B C : PTm) :
+  Lemma commutativity0 (A B C : PTm) :
     R A B ->
     (RERed.R A C ->
     exists D, RERed.R B D /\ R C D) /\
@@ -3333,26 +3339,27 @@ Module Sub1.
     exists D, RERed.R A D /\ R D C).
   Proof.
     move => h. move : C.
-    elim : n A B / h; hauto lq:on ctrs:RERed.R, R inv:RERed.R.
+    elim : A B / h; hauto lq:on ctrs:RERed.R, R inv:RERed.R.
   Qed.
 
-  Lemma commutativity_Ls n (A B C : PTm) :
+  Lemma commutativity_Ls (A B C : PTm) :
     R A B ->
     rtc RERed.R A C ->
     exists D, rtc RERed.R B D /\ R C D.
   Proof.
-    move => + h. move : B. induction h; sauto lq:on use:commutativity0.
+    move => + h. move : B.
+    induction h; ecrush use:commutativity0.
   Qed.
 
-  Lemma commutativity_Rs n (A B C : PTm) :
+  Lemma commutativity_Rs (A B C : PTm) :
     R A B ->
     rtc RERed.R B C ->
     exists D, rtc RERed.R A D /\ R D C.
   Proof.
-    move => + h. move : A. induction h; sauto lq:on use:commutativity0.
+    move => + h. move : A. induction h; ecrush use:commutativity0.
   Qed.
 
-  Lemma sn_preservation  : forall n,
+  Lemma sn_preservation  :
   (forall (a : PTm) (s : SNe a), forall b, R a b \/ R b a -> a = b) /\
   (forall (a : PTm) (s : SN a), forall b, R a b \/ R b a -> SN b) /\
   (forall (a b : PTm) (_ : TRedSN a b), forall c, R a c \/ R c a -> a = c).
@@ -3360,30 +3367,30 @@ Module Sub1.
     apply sn_mutual; hauto lq:on inv:R ctrs:SN.
   Qed.
 
-  Lemma bind_preservation n (a b : PTm) :
+  Lemma bind_preservation (a b : PTm) :
     R a b -> isbind a = isbind b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma univ_preservation n (a b : PTm) :
+  Lemma univ_preservation (a b : PTm) :
     R a b -> isuniv a = isuniv b.
   Proof. hauto q:on inv:R. Qed.
 
-  Lemma sne_preservation n (a b : PTm) :
+  Lemma sne_preservation (a b : PTm) :
     R a b -> SNe a <-> SNe b.
   Proof. hfcrush use:sn_preservation. Qed.
 
-  Lemma renaming n m (a b : PTm) (ξ : fin n -> fin m) :
-    R a b -> R (ren_PTm a) (ren_PTm b).
+  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; hauto lq:on ctrs:R.
+    move => h. move : ξ.
+    elim : a b /h; hauto lq:on ctrs:R.
   Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    R a b -> R (subst_PTm a) (subst_PTm b).
-  Proof. move => h. move : m ρ. elim : n a b /h; hauto lq:on ctrs:R. Qed.
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Proof. move => h. move : ρ. elim : a b /h; hauto lq:on ctrs:R. Qed.
 
-  Lemma hne_refl n (a b : PTm) :
+  Lemma hne_refl (a b : PTm) :
     ishne a \/ ishne b -> R a b -> a = b.
   Proof. hauto q:on inv:R. Qed.
 
@@ -3392,7 +3399,7 @@ End Sub1.
 Module ESub.
   Definition R (a b : PTm) := exists c0 c1, rtc ERed.R a c0 /\ rtc ERed.R b c1 /\ Sub1.R c0 c1.
 
-  Lemma pi_inj n (A0 A1 : PTm) B0 B1 :
+  Lemma pi_inj (A0 A1 : PTm) B0 B1 :
     R (PBind PPi A0 B0) (PBind PPi A1 B1) ->
     R A1 A0 /\ R B0 B1.
   Proof.
@@ -3402,7 +3409,7 @@ Module ESub.
     sauto lq:on rew:off inv:Sub1.R.
   Qed.
 
-  Lemma sig_inj n (A0 A1 : PTm) B0 B1 :
+  Lemma sig_inj (A0 A1 : PTm) B0 B1 :
     R (PBind PSig A0 B0) (PBind PSig A1 B1) ->
     R A0 A1 /\ R B0 B1.
   Proof.
@@ -3412,7 +3419,7 @@ Module ESub.
     sauto lq:on rew:off inv:Sub1.R.
   Qed.
 
-  Lemma suc_inj n (a b : PTm) :
+  Lemma suc_inj (a b : PTm) :
     R (PSuc a) (PSuc b) ->
     R a b.
   Proof.
@@ -3424,10 +3431,10 @@ End ESub.
 Module Sub.
   Definition R (a b : PTm) := exists c d, rtc RERed.R a c /\ rtc RERed.R b d /\ Sub1.R c d.
 
-  Lemma refl n (a : PTm) : R a a.
+  Lemma refl (a : PTm) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
-  Lemma ToJoin n (a b : PTm) :
+  Lemma ToJoin (a b : PTm) :
     ishne a \/ ishne b ->
     R a b ->
     DJoin.R a b.
@@ -3437,7 +3444,7 @@ Module Sub.
     hauto lq:on rew:off use:Sub1.hne_refl.
   Qed.
 
-  Lemma transitive n (a b c : PTm) : SN b -> R a b -> R b c -> R a c.
+  Lemma transitive (a b c : PTm) : SN b -> R a b -> R b c -> R a c.
   Proof.
     rewrite /R.
     move => h  [a0][b0][ha][hb]ha0b0 [b1][c0][hb'][hc]hb1c0.
@@ -3449,10 +3456,10 @@ Module Sub.
     exists a',c'; hauto l:on use:Sub1.transitive, @relations.rtc_transitive.
   Qed.
 
-  Lemma FromJoin n (a b : PTm) : DJoin.R a b -> R a b.
+  Lemma FromJoin (a b : PTm) : DJoin.R a b -> R a b.
   Proof. sfirstorder. Qed.
 
-  Lemma PiCong n (A0 A1 : PTm) B0 B1 :
+  Lemma PiCong (A0 A1 : PTm) B0 B1 :
     R A1 A0 ->
     R B0 B1 ->
     R (PBind PPi A0 B0) (PBind PPi A1 B1).
@@ -3464,7 +3471,7 @@ Module Sub.
     repeat split; eauto using REReds.BindCong, Sub1.PiCong.
   Qed.
 
-  Lemma SigCong n (A0 A1 : PTm) B0 B1 :
+  Lemma SigCong (A0 A1 : PTm) B0 B1 :
     R A0 A1 ->
     R B0 B1 ->
     R (PBind PSig A0 B0) (PBind PSig A1 B1).
@@ -3476,12 +3483,12 @@ Module Sub.
     repeat split; eauto using REReds.BindCong, Sub1.SigCong.
   Qed.
 
-  Lemma UnivCong n i j :
+  Lemma UnivCong i j :
     i <= j ->
-    @R n (PUniv i) (PUniv j).
+    @R (PUniv i) (PUniv j).
   Proof. hauto lq:on ctrs:Sub1.R, rtc. Qed.
 
-  Lemma sne_nat_noconf n (a b : PTm) :
+  Lemma sne_nat_noconf (a b : PTm) :
     R a b -> SNe a -> isnat b -> False.
   Proof.
     move => [c [d [h0 [h1 h2]]]] *.
@@ -3489,7 +3496,7 @@ Module Sub.
     move : h2. clear. hauto q:on inv:Sub1.R, SNe.
   Qed.
 
-  Lemma nat_sne_noconf n (a b : PTm) :
+  Lemma nat_sne_noconf (a b : PTm) :
     R a b -> isnat a -> SNe b -> False.
   Proof.
     move => [c [d [h0 [h1 h2]]]] *.
@@ -3497,7 +3504,7 @@ Module Sub.
     move : h2. clear. hauto q:on inv:Sub1.R, SNe.
   Qed.
 
-  Lemma sne_bind_noconf n (a b : PTm) :
+  Lemma sne_bind_noconf (a b : PTm) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
     rewrite /R.
@@ -3507,7 +3514,7 @@ Module Sub.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma hne_bind_noconf n (a b : PTm) :
+  Lemma hne_bind_noconf (a b : PTm) :
     R a b -> ishne a -> isbind b -> False.
   Proof.
     rewrite /R.
@@ -3518,7 +3525,7 @@ Module Sub.
     clear. case : d => //=.
   Qed.
 
-  Lemma bind_sne_noconf n (a b : PTm) :
+  Lemma bind_sne_noconf (a b : PTm) :
     R a b -> SNe b -> isbind a -> False.
   Proof.
     rewrite /R.
@@ -3528,14 +3535,14 @@ Module Sub.
     qauto l:on inv:SNe.
   Qed.
 
-  Lemma sne_univ_noconf n (a b : PTm) :
+  Lemma sne_univ_noconf (a b : PTm) :
     R a b -> SNe a -> isuniv b -> False.
   Proof.
     hauto l:on use:REReds.sne_preservation,
           REReds.univ_preservation, Sub1.sne_preservation, Sub1.univ_preservation inv:SNe.
   Qed.
 
-  Lemma univ_sne_noconf n (a b : PTm) :
+  Lemma univ_sne_noconf (a b : PTm) :
     R a b -> SNe b -> isuniv a -> False.
   Proof.
     move => [c[d] [? []]] *.
@@ -3545,7 +3552,7 @@ Module Sub.
     clear. case : c => //=. inversion 2.
   Qed.
 
-  Lemma univ_nat_noconf n (a b : PTm)  :
+  Lemma univ_nat_noconf (a b : PTm)  :
     R a b -> isuniv b -> isnat a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3555,7 +3562,7 @@ Module Sub.
     clear. case : d => //=.
   Qed.
 
-  Lemma nat_univ_noconf n (a b : PTm)  :
+  Lemma nat_univ_noconf (a b : PTm)  :
     R a b -> isnat b -> isuniv a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3565,7 +3572,7 @@ Module Sub.
     clear. case : d => //=.
   Qed.
 
-  Lemma bind_nat_noconf n (a b : PTm)  :
+  Lemma bind_nat_noconf (a b : PTm)  :
     R a b -> isbind b -> isnat a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3576,7 +3583,7 @@ Module Sub.
     case : d h1 => //=.
   Qed.
 
-  Lemma nat_bind_noconf n (a b : PTm)  :
+  Lemma nat_bind_noconf (a b : PTm)  :
     R a b -> isnat b -> isbind a -> False.
   Proof.
     move => [c[d] [? []]] h0 h1 h2 h3.
@@ -3587,7 +3594,7 @@ Module Sub.
     case : d h1 => //=.
   Qed.
 
-  Lemma bind_univ_noconf n (a b : PTm) :
+  Lemma bind_univ_noconf (a b : PTm) :
     R a b -> isbind a -> isuniv b -> False.
   Proof.
     move => [c[d] [h0 [h1 h1']]] h2 h3.
@@ -3598,7 +3605,7 @@ Module Sub.
     case : c => //=; itauto.
   Qed.
 
-  Lemma univ_bind_noconf n (a b : PTm) :
+  Lemma univ_bind_noconf (a b : PTm) :
     R a b -> isbind b -> isuniv a -> False.
   Proof.
     move => [c[d] [h0 [h1 h1']]] h2 h3.
@@ -3609,7 +3616,7 @@ Module Sub.
     case : c => //=; itauto.
   Qed.
 
-  Lemma bind_inj n p0 p1 (A0 A1 : PTm) B0 B1 :
+  Lemma bind_inj p0 p1 (A0 A1 : PTm) B0 B1 :
     R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ (if p0 is PPi then R A1 A0 else R A0 A1) /\ R B0 B1.
   Proof.
@@ -3620,13 +3627,13 @@ Module Sub.
     inversion h; subst; sauto lq:on.
   Qed.
 
-  Lemma univ_inj n i j :
-    @R n (PUniv i) (PUniv j)  -> i <= j.
+  Lemma univ_inj i j :
+    @R (PUniv i) (PUniv j)  -> i <= j.
   Proof.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
-  Lemma suc_inj n  (A0 A1 : PTm)  :
+  Lemma suc_inj (A0 A1 : PTm)  :
     R (PSuc A0) (PSuc A1) ->
     R A0 A1.
   Proof.
@@ -3634,7 +3641,7 @@ Module Sub.
   Qed.
 
 
-  Lemma cong n m (a b : PTm (S n)) c d (ρ : fin n -> PTm) :
+  Lemma cong (a b : PTm) c d (ρ : nat -> PTm) :
     R a b -> DJoin.R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) b).
   Proof.
     rewrite /R.
@@ -3642,23 +3649,23 @@ Module Sub.
     move => [cd][h3]h4.
     exists (subst_PTm (scons cd ρ) a0), (subst_PTm (scons cd ρ) b0).
     repeat split.
-    hauto l:on use:REReds.cong' inv:option.
-    hauto l:on use:REReds.cong' inv:option.
+    hauto l:on use:REReds.cong' inv:nat.
+    hauto l:on use:REReds.cong' inv:nat.
     eauto using Sub1.substing.
   Qed.
 
-  Lemma substing n m (a b : PTm) (ρ : fin n -> PTm) :
-    R a b -> R (subst_PTm a) (subst_PTm b).
+  Lemma substing (a b : PTm) (ρ : nat -> PTm) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     rewrite /R.
     move => [a0][b0][h0][h1]h2.
     hauto ctrs:rtc use:REReds.cong', Sub1.substing.
   Qed.
 
-  Lemma ToESub n (a b : PTm) : nf a -> nf b -> R a b -> ESub.R a b.
+  Lemma ToESub (a b : PTm) : nf a -> nf b -> R a b -> ESub.R a b.
   Proof. hauto q:on use:REReds.ToEReds. Qed.
 
-  Lemma standardization n (a b : PTm) :
+  Lemma standardization (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc RRed.R a va /\ rtc RRed.R b vb /\ nf va /\ nf vb /\ ESub.R va vb.
   Proof.
@@ -3676,7 +3683,7 @@ Module Sub.
     hauto lq:on.
   Qed.
 
-  Lemma standardization_lo n (a b : PTm) :
+  Lemma standardization_lo (a b : PTm) :
     SN a -> SN b -> R a b ->
     exists va vb, rtc LoRed.R a va /\ rtc LoRed.R b vb /\ nf va /\ nf vb /\ ESub.R va vb.
   Proof.

From 47e21df801b8e43967ab4becf7e7aad1258ceaf0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 3 Mar 2025 01:15:21 -0500
Subject: [PATCH 05/24] Finish refactoring logical relations

---
 theories/common.v |   2 +
 theories/logrel.v | 692 +++++++++++++++++++++-------------------------
 2 files changed, 318 insertions(+), 376 deletions(-)

diff --git a/theories/common.v b/theories/common.v
index 79b0b19..a890edd 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -10,6 +10,8 @@ Inductive lookup : nat -> list PTm -> PTm -> Prop :=
   lookup i Γ A ->
   lookup (S i) (cons B Γ) (ren_PTm shift A).
 
+Derive Inversion lookup_inv with (forall i Γ A, lookup i Γ A).
+
 Definition renaming_ok (Γ : list PTm) (Δ : list PTm) (ξ : nat -> nat) :=
   forall i A, lookup i Δ A -> lookup (ξ i) Γ (ren_PTm ξ A).
 
diff --git a/theories/logrel.v b/theories/logrel.v
index a245362..a113437 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax.
 Require Import common fp_red.
 From Hammer Require Import Tactics.
 From Equations Require Import Equations.
@@ -8,19 +8,19 @@ From stdpp Require Import relations (rtc(..), rtc_subrel).
 Import Psatz.
 Require Import Cdcl.Itauto.
 
-Definition ProdSpace {n} (PA : PTm n -> Prop)
-  (PF : PTm n -> (PTm n -> Prop) -> Prop) b : Prop :=
+Definition ProdSpace (PA : PTm -> Prop)
+  (PF : PTm -> (PTm -> Prop) -> Prop) b : Prop :=
   forall a PB, PA a -> PF a PB -> PB (PApp b a).
 
-Definition SumSpace {n} (PA : PTm n -> Prop)
-  (PF : PTm n -> (PTm n -> Prop) -> Prop) t : Prop :=
+Definition SumSpace (PA : PTm -> Prop)
+  (PF : PTm -> (PTm -> Prop) -> Prop) t : Prop :=
   (exists v, rtc TRedSN t v /\ SNe v) \/ exists a b, rtc TRedSN t (PPair a b) /\ PA a /\ (forall PB, PF a PB -> PB b).
 
-Definition BindSpace {n} p := if p is PPi then @ProdSpace n else SumSpace.
+Definition BindSpace p := if p is PPi then ProdSpace else SumSpace.
 
 Reserved Notation "⟦ A ⟧ i ;; I ↘ S" (at level 70).
 
-Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop) -> Prop :=
+Inductive InterpExt i (I : nat -> PTm -> Prop) : PTm -> (PTm -> Prop) -> Prop :=
 | InterpExt_Ne A :
   SNe A ->
   ⟦ A ⟧ i ;; I ↘ (fun a => exists v, rtc TRedSN a v /\ SNe v)
@@ -44,7 +44,7 @@ Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop)
   ⟦ A ⟧ i ;; I ↘ PA
 where "⟦ A ⟧ i ;; I ↘ S" := (InterpExt i I A S).
 
-Lemma InterpExt_Univ' n i  I j (PF : PTm n -> Prop) :
+Lemma InterpExt_Univ' i  I j (PF : PTm -> Prop) :
   PF = I j ->
   j < i ->
   ⟦ PUniv j ⟧ i ;; I ↘ PF.
@@ -52,16 +52,15 @@ Proof. hauto lq:on ctrs:InterpExt. Qed.
 
 Infix "<?" := Compare_dec.lt_dec (at level 60).
 
-Equations InterpUnivN n (i : nat) : PTm n -> (PTm n -> Prop) -> Prop by wf i lt :=
-  InterpUnivN n i := @InterpExt n i
+Equations InterpUnivN (i : nat) : PTm -> (PTm -> Prop) -> Prop by wf i lt :=
+  InterpUnivN i := @InterpExt i
                      (fun j A =>
                         match j <? i  with
-                        | left _ => exists PA, InterpUnivN n j A PA
+                        | left _ => exists PA, InterpUnivN j A PA
                         | right _ => False
                         end).
-Arguments InterpUnivN {n}.
 
-Lemma InterpExt_lt_impl n i I I' A (PA : PTm n -> Prop) :
+Lemma InterpExt_lt_impl i I I' A (PA : PTm -> Prop) :
   (forall j, j < i -> I j = I' j) ->
   ⟦ A ⟧ i ;; I ↘ PA ->
   ⟦ A ⟧ i ;; I' ↘ PA.
@@ -75,7 +74,7 @@ Proof.
   - hauto lq:on ctrs:InterpExt.
 Qed.
 
-Lemma InterpExt_lt_eq n i I I' A (PA : PTm n -> Prop) :
+Lemma InterpExt_lt_eq i I I' A (PA : PTm -> Prop) :
   (forall j, j < i -> I j = I' j) ->
   ⟦ A ⟧ i ;; I ↘ PA =
   ⟦ A ⟧ i ;; I' ↘ PA.
@@ -87,8 +86,8 @@ Qed.
 
 Notation "⟦ A ⟧ i ↘ S" := (InterpUnivN i A S) (at level 70).
 
-Lemma InterpUnivN_nolt n i :
-  @InterpUnivN n i = @InterpExt n i (fun j (A : PTm n) => exists PA, ⟦ A ⟧ j ↘ PA).
+Lemma InterpUnivN_nolt i :
+  @InterpUnivN i = @InterpExt i (fun j (A : PTm ) => exists PA, ⟦ A ⟧ j ↘ PA).
 Proof.
   simp InterpUnivN.
   extensionality A. extensionality PA.
@@ -98,64 +97,62 @@ Qed.
 
 #[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv.
 
-Check InterpExt_ind.
-
 Lemma InterpUniv_ind
-  : forall n (P : nat -> PTm n -> (PTm n -> Prop) -> Prop),
-       (forall i (A : PTm n), SNe A -> P i A (fun a : PTm n => exists v : PTm n, rtc TRedSN a v /\ SNe v)) ->
-       (forall i (p : BTag) (A : PTm n) (B : PTm (S n)) (PA : PTm n -> Prop)
-          (PF : PTm n -> (PTm n -> Prop) -> Prop),
+  : forall (P : nat -> PTm -> (PTm -> Prop) -> Prop),
+       (forall i (A : PTm), SNe A -> P i A (fun a : PTm => exists v : PTm , rtc TRedSN a v /\ SNe v)) ->
+       (forall i (p : BTag) (A : PTm ) (B : PTm ) (PA : PTm  -> Prop)
+          (PF : PTm  -> (PTm  -> Prop) -> Prop),
         ⟦ A ⟧ i ↘ PA ->
         P i A PA ->
-        (forall a : PTm n, PA a -> exists PB : PTm n -> Prop, PF a PB) ->
-        (forall (a : PTm n) (PB : PTm n -> Prop), PF a PB -> ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB) ->
-        (forall (a : PTm n) (PB : PTm n -> Prop), PF a PB -> P i (subst_PTm (scons a VarPTm) B) PB) ->
+        (forall a : PTm , PA a -> exists PB : PTm  -> Prop, PF a PB) ->
+        (forall (a : PTm ) (PB : PTm  -> Prop), PF a PB -> ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB) ->
+        (forall (a : PTm ) (PB : PTm  -> Prop), PF a PB -> P i (subst_PTm (scons a VarPTm) B) PB) ->
         P i (PBind p A B) (BindSpace p PA PF)) ->
        (forall i, P i PNat SNat) ->
        (forall i j : nat, j < i ->  (forall A PA, ⟦ A ⟧ j ↘ PA -> P j A PA) -> P i (PUniv j) (fun A => exists PA, ⟦ A ⟧ j ↘ PA)) ->
-       (forall i (A A0 : PTm n) (PA : PTm n -> Prop), TRedSN A A0 -> ⟦ A0 ⟧ i ↘ PA -> P i A0 PA -> P i A PA) ->
-       forall i (p : PTm n) (P0 : PTm n -> Prop), ⟦ p ⟧ i ↘ P0 -> P i p P0.
+       (forall i (A A0 : PTm ) (PA : PTm  -> Prop), TRedSN A A0 -> ⟦ A0 ⟧ i ↘ PA -> P i A0 PA -> P i A PA) ->
+       forall i (p : PTm ) (P0 : PTm  -> Prop), ⟦ p ⟧ i ↘ P0 -> P i p P0.
 Proof.
-  move => n P  hSN hBind hNat hUniv hRed.
+  move => P  hSN hBind hNat hUniv hRed.
   elim /Wf_nat.lt_wf_ind => i ih . simp InterpUniv.
   move => A PA. move => h. set I := fun _ => _ in h.
   elim : A PA / h; rewrite -?InterpUnivN_nolt; eauto.
 Qed.
 
-Derive Dependent Inversion iinv with (forall n i I (A : PTm n) PA, InterpExt i I A PA) Sort Prop.
+Derive Dependent Inversion iinv with (forall i I (A : PTm ) PA, InterpExt i I A PA) Sort Prop.
 
-Lemma InterpUniv_Ne n i (A : PTm n) :
+Lemma InterpUniv_Ne i (A : PTm) :
   SNe A ->
   ⟦ A ⟧ i ↘ (fun a => exists v, rtc TRedSN a v /\ SNe v).
 Proof. simp InterpUniv. apply InterpExt_Ne. Qed.
 
-Lemma InterpUniv_Bind n i p A B PA PF :
-  ⟦ A : PTm n ⟧ i ↘ PA ->
+Lemma InterpUniv_Bind  i p A B PA PF :
+  ⟦ A ⟧ i ↘ PA ->
   (forall a, PA a -> exists PB, PF a PB) ->
   (forall a PB, PF a PB -> ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB) ->
   ⟦ PBind p A B ⟧ i ↘ BindSpace p PA PF.
 Proof. simp InterpUniv. apply InterpExt_Bind. Qed.
 
-Lemma InterpUniv_Univ n i j :
-  j < i -> ⟦ PUniv j : PTm n ⟧ i ↘ (fun A => exists PA, ⟦ A ⟧ j ↘ PA).
+Lemma InterpUniv_Univ i j :
+  j < i -> ⟦ PUniv j ⟧ i ↘ (fun A => exists PA, ⟦ A ⟧ j ↘ PA).
 Proof.
   simp InterpUniv. simpl.
   apply InterpExt_Univ'. by simp InterpUniv.
 Qed.
 
-Lemma InterpUniv_Step i n A A0 PA :
+Lemma InterpUniv_Step i A A0 PA :
   TRedSN A A0 ->
-  ⟦ A0 : PTm n ⟧ i ↘ PA ->
+  ⟦ A0 ⟧ i ↘ PA ->
   ⟦ A ⟧ i ↘ PA.
 Proof. simp InterpUniv. apply InterpExt_Step. Qed.
 
-Lemma InterpUniv_Nat i n :
-  ⟦ PNat : PTm n ⟧ i ↘ SNat.
+Lemma InterpUniv_Nat i :
+  ⟦ PNat  ⟧ i ↘ SNat.
 Proof. simp InterpUniv. apply InterpExt_Nat. Qed.
 
 #[export]Hint Resolve InterpUniv_Bind InterpUniv_Step InterpUniv_Ne InterpUniv_Univ : InterpUniv.
 
-Lemma InterpExt_cumulative n i j I (A : PTm n) PA :
+Lemma InterpExt_cumulative i j I (A : PTm ) PA :
   i <= j ->
    ⟦ A ⟧ i ;; I ↘ PA ->
    ⟦ A ⟧ j ;; I ↘ PA.
@@ -165,18 +162,18 @@ Proof.
     hauto l:on ctrs:InterpExt solve+:(by lia).
 Qed.
 
-Lemma InterpUniv_cumulative n i (A : PTm n) PA :
+Lemma InterpUniv_cumulative i (A : PTm) PA :
    ⟦ A ⟧ i ↘ PA -> forall j, i <= j ->
    ⟦ A ⟧ j ↘ PA.
 Proof.
   hauto l:on rew:db:InterpUniv use:InterpExt_cumulative.
 Qed.
 
-Definition CR {n} (P : PTm n -> Prop) :=
+Definition CR (P : PTm -> Prop) :=
   (forall a, P a -> SN a) /\
     (forall a, SNe a -> P a).
 
-Lemma N_Exps n (a b : PTm n) :
+Lemma N_Exps (a b : PTm) :
   rtc TRedSN a b ->
   SN b ->
   SN a.
@@ -184,21 +181,22 @@ Proof.
   induction 1; eauto using N_Exp.
 Qed.
 
-Lemma CR_SNat : forall n,  @CR n SNat.
+Lemma CR_SNat : CR SNat.
 Proof.
-  move => n. rewrite /CR.
+  rewrite /CR.
   split.
   induction 1; hauto q:on ctrs:SN,SNe.
   hauto lq:on ctrs:SNat.
 Qed.
 
-Lemma adequacy : forall i n A PA,
-   ⟦ A : PTm n ⟧ i ↘ PA ->
+Lemma adequacy : forall i A PA,
+   ⟦ A ⟧ i ↘ PA ->
   CR PA /\ SN A.
 Proof.
-  move => + n. apply : InterpUniv_ind.
+  apply : InterpUniv_ind.
   - hauto l:on use:N_Exps ctrs:SN,SNe.
   - move => i p A B PA PF hPA [ihA0 ihA1] hTot hRes ihPF.
+    set PBot := (VarPTm var_zero).
     have hb : PA PBot by hauto q:on ctrs:SNe.
     have hb' : SN PBot by hauto q:on ctrs:SN, SNe.
     rewrite /CR.
@@ -218,18 +216,18 @@ Proof.
       apply : N_Exp; eauto using N_β.
       hauto lq:on.
       qauto l:on use:SN_AppInv, SN_NoForbid.P_AbsInv.
-  - qauto use:CR_SNat.
+  - sfirstorder use:CR_SNat.
   - hauto l:on ctrs:InterpExt rew:db:InterpUniv.
   - hauto l:on ctrs:SN unfold:CR.
 Qed.
 
-Lemma InterpUniv_Steps i n A A0 PA :
+Lemma InterpUniv_Steps i A A0 PA :
   rtc TRedSN A A0 ->
-  ⟦ A0 : PTm n ⟧ i ↘ PA ->
+  ⟦ A0 ⟧ i ↘ PA ->
   ⟦ A ⟧ i ↘ PA.
 Proof. induction 1; hauto l:on use:InterpUniv_Step. Qed.
 
-Lemma InterpUniv_back_clos n i (A : PTm n) PA :
+Lemma InterpUniv_back_clos i (A : PTm ) PA :
     ⟦ A ⟧ i ↘ PA ->
     forall a b, TRedSN a b ->
            PA b -> PA a.
@@ -248,7 +246,7 @@ Proof.
   - hauto l:on use:InterpUniv_Step.
 Qed.
 
-Lemma InterpUniv_back_closs n i (A : PTm n) PA :
+Lemma InterpUniv_back_closs i (A : PTm) PA :
     ⟦ A ⟧ i ↘ PA ->
     forall a b, rtc TRedSN a b ->
            PA b -> PA a.
@@ -256,7 +254,7 @@ Proof.
   induction 2; hauto lq:on ctrs:rtc use:InterpUniv_back_clos.
 Qed.
 
-Lemma InterpUniv_case n i (A : PTm n) PA :
+Lemma InterpUniv_case i (A : PTm) PA :
   ⟦ A ⟧ i ↘ PA ->
   exists H, rtc TRedSN A H /\  ⟦ H ⟧ i ↘ PA /\ (SNe H \/ isbind H \/ isuniv H \/ isnat H).
 Proof.
@@ -268,21 +266,21 @@ Proof.
   hauto lq:on ctrs:rtc.
 Qed.
 
-Lemma redsn_preservation_mutual n :
-  (forall (a : PTm n) (s : SNe a), forall b, TRedSN a b -> False) /\
-    (forall (a : PTm n) (s : SN a), forall b, TRedSN a b -> SN b) /\
-  (forall (a b : PTm n) (_ : TRedSN a b), forall c, TRedSN a c -> b = c).
+Lemma redsn_preservation_mutual :
+  (forall (a : PTm) (s : SNe a), forall b, TRedSN a b -> False) /\
+    (forall (a : PTm) (s : SN a), forall b, TRedSN a b -> SN b) /\
+  (forall (a b : PTm) (_ : TRedSN a b), forall c, TRedSN a c -> b = c).
 Proof.
-  move : n. apply sn_mutual; sauto lq:on rew:off.
+  apply sn_mutual; sauto lq:on rew:off.
 Qed.
 
-Lemma redsns_preservation : forall n a b, @SN n a -> rtc TRedSN a b -> SN b.
+Lemma redsns_preservation : forall a b, SN a -> rtc TRedSN a b -> SN b.
 Proof. induction 2; sfirstorder use:redsn_preservation_mutual ctrs:rtc. Qed.
 
 #[export]Hint Resolve Sub.sne_bind_noconf Sub.sne_univ_noconf Sub.bind_univ_noconf
   Sub.bind_sne_noconf Sub.univ_sne_noconf Sub.univ_bind_noconf Sub.nat_bind_noconf Sub.bind_nat_noconf Sub.sne_nat_noconf Sub.nat_sne_noconf Sub.univ_nat_noconf Sub.nat_univ_noconf: noconf.
 
-Lemma InterpUniv_SNe_inv n i (A : PTm n) PA :
+Lemma InterpUniv_SNe_inv i (A : PTm) PA :
   SNe A ->
   ⟦ A ⟧ i ↘ PA ->
   PA = (fun a => exists v, rtc TRedSN a v /\ SNe v).
@@ -291,9 +289,9 @@ Proof.
   hauto lq:on rew:off inv:InterpExt,SNe use:redsn_preservation_mutual.
 Qed.
 
-Lemma InterpUniv_Bind_inv n i p A B S :
+Lemma InterpUniv_Bind_inv i p A B S :
   ⟦ PBind p A B ⟧ i ↘ S -> exists PA PF,
-  ⟦ A : PTm n ⟧ i ↘ PA /\
+  ⟦ A ⟧ i ↘ PA /\
   (forall a, PA a -> exists PB, PF a PB) /\
   (forall a PB, PF a PB -> ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB) /\
   S = BindSpace p PA PF.
@@ -303,16 +301,16 @@ Proof. simp InterpUniv.
        sauto lq:on.
 Qed.
 
-Lemma InterpUniv_Nat_inv n i S :
-  ⟦ PNat : PTm n ⟧ i ↘ S -> S = SNat.
+Lemma InterpUniv_Nat_inv i S :
+  ⟦ PNat  ⟧ i ↘ S -> S = SNat.
 Proof.
   simp InterpUniv.
-       inversion 1; try hauto inv:SNe q:on use:redsn_preservation_mutual.
-       sauto lq:on.
+  inversion 1; try hauto inv:SNe q:on use:redsn_preservation_mutual.
+  sauto lq:on.
 Qed.
 
-Lemma InterpUniv_Univ_inv n i j S :
-  ⟦ PUniv j : PTm n ⟧ i ↘ S ->
+Lemma InterpUniv_Univ_inv i j S :
+  ⟦ PUniv j ⟧ i ↘ S ->
   S = (fun A => exists PA, ⟦ A ⟧ j ↘ PA) /\ j < i.
 Proof.
   simp InterpUniv. inversion 1;
@@ -321,9 +319,9 @@ Proof.
   subst. hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma bindspace_impl n p (PA PA0 : PTm n -> Prop) PF PF0 b  :
+Lemma bindspace_impl p (PA PA0 : PTm -> Prop) PF PF0 b  :
   (forall x, if p is PPi then (PA0 x -> PA x) else (PA x -> PA0 x)) ->
-  (forall (a : PTm n) (PB PB0 : PTm n -> Prop), PA0 a -> PF a PB -> PF0 a PB0 -> (forall x, PB x ->  PB0 x)) ->
+  (forall (a : PTm ) (PB PB0 : PTm  -> Prop), PA0 a -> PF a PB -> PF0 a PB0 -> (forall x, PB x ->  PB0 x)) ->
   (forall a, PA a -> exists PB, PF a PB) ->
   (forall a, PA0 a -> exists PB0, PF0 a PB0) ->
   (BindSpace p PA PF b -> BindSpace p PA0 PF0 b).
@@ -341,7 +339,7 @@ Proof.
     hauto lq:on.
 Qed.
 
-Lemma InterpUniv_Sub' n i (A B : PTm n) PA PB :
+Lemma InterpUniv_Sub' i (A B : PTm) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
   ((Sub.R A B -> forall x, PA x -> PB x) /\
@@ -416,7 +414,7 @@ Proof.
       have {h}{}hAB : Sub.R PNat H by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have {}h0 : isnat H.
       suff : ~ isbind H /\ ~ isuniv H /\ ~ SNe H by itauto.
-      have : @isnat n PNat by simpl.
+      have : @isnat PNat by simpl.
       move : h0 hAB. clear. qauto l:on db:noconf.
       case : H h1 hAB h0 => //=.
       move /InterpUniv_Nat_inv. scongruence.
@@ -427,7 +425,7 @@ Proof.
       have {h}{}hAB : Sub.R H PNat by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have {}h0 : isnat H.
       suff : ~ isbind H /\ ~ isuniv H /\ ~ SNe H by itauto.
-      have : @isnat n PNat by simpl.
+      have : @isnat PNat by simpl.
       move : h0 hAB. clear. qauto l:on db:noconf.
       case : H h1 hAB h0 => //=.
       move /InterpUniv_Nat_inv. scongruence.
@@ -468,7 +466,7 @@ Proof.
       qauto l:on.
 Qed.
 
-Lemma InterpUniv_Sub0 n i (A B : PTm n) PA PB :
+Lemma InterpUniv_Sub0 i (A B : PTm) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
   Sub.R A B -> forall x, PA x -> PB x.
@@ -477,7 +475,7 @@ Proof.
   move => [+ _]. apply.
 Qed.
 
-Lemma InterpUniv_Sub n i j (A B : PTm n) PA PB :
+Lemma InterpUniv_Sub i j (A B : PTm) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ j ↘ PB ->
   Sub.R A B -> forall x, PA x -> PB x.
@@ -490,7 +488,7 @@ Proof.
   apply.
 Qed.
 
-Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
+Lemma InterpUniv_Join i (A B : PTm) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
   DJoin.R A B ->
@@ -504,13 +502,13 @@ Proof.
   firstorder.
 Qed.
 
-Lemma InterpUniv_Functional n i  (A : PTm n) PA PB :
+Lemma InterpUniv_Functional i  (A : PTm) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ A ⟧ i ↘ PB ->
   PA = PB.
 Proof. hauto l:on use:InterpUniv_Join, DJoin.refl. Qed.
 
-Lemma InterpUniv_Join' n i j (A B : PTm n) PA PB :
+Lemma InterpUniv_Join' i j (A B : PTm) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ j ↘ PB ->
   DJoin.R A B ->
@@ -523,16 +521,16 @@ Proof.
   eauto using InterpUniv_Join.
 Qed.
 
-Lemma InterpUniv_Functional' n i j A PA PB :
-  ⟦ A : PTm n ⟧ i ↘ PA ->
+Lemma InterpUniv_Functional' i j A PA PB :
+  ⟦ A ⟧ i ↘ PA ->
   ⟦ A ⟧ j ↘ PB ->
   PA = PB.
 Proof.
   hauto l:on use:InterpUniv_Join', DJoin.refl.
 Qed.
 
-Lemma InterpUniv_Bind_inv_nopf n i p A B P (h :  ⟦PBind p A B ⟧ i ↘ P) :
-  exists (PA : PTm n -> Prop),
+Lemma InterpUniv_Bind_inv_nopf i p A B P (h :  ⟦PBind p A B ⟧ i ↘ P) :
+  exists (PA : PTm -> Prop),
      ⟦ A ⟧ i ↘ PA /\
     (forall a, PA a -> exists PB, ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB) /\
       P = BindSpace p PA (fun a PB => ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB).
@@ -559,19 +557,20 @@ Proof.
       split; hauto q:on use:InterpUniv_Functional.
 Qed.
 
-Definition ρ_ok {n} (Γ : fin n -> PTm n) (ρ : fin n -> PTm 0) := forall i k PA,
-    ⟦ subst_PTm ρ (Γ i) ⟧ k ↘ PA -> PA (ρ i).
+Definition ρ_ok (Γ : list PTm) (ρ : nat -> PTm) := forall i k A PA,
+    lookup i Γ A ->
+    ⟦ subst_PTm ρ A ⟧ k ↘ PA -> PA (ρ i).
 
-Definition SemWt {n} Γ (a A : PTm n) := forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a).
+Definition SemWt Γ (a A : PTm) := forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a).
 Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70).
 
-Definition SemEq {n} Γ (a b A : PTm n) := DJoin.R a b /\ forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a) /\ PA (subst_PTm ρ b).
+Definition SemEq Γ (a b A : PTm) := DJoin.R a b /\ forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a) /\ PA (subst_PTm ρ b).
 Notation "Γ ⊨ a ≡ b ∈ A" := (SemEq Γ a b A) (at level 70).
 
-Definition SemLEq {n} Γ (A B : PTm n) := Sub.R A B /\ exists i, forall ρ, ρ_ok Γ ρ -> exists S0 S1, ⟦ subst_PTm ρ A ⟧ i ↘ S0 /\ ⟦ subst_PTm ρ B ⟧ i ↘ S1.
+Definition SemLEq Γ (A B : PTm) := Sub.R A B /\ exists i, forall ρ, ρ_ok Γ ρ -> exists S0 S1, ⟦ subst_PTm ρ A ⟧ i ↘ S0 /\ ⟦ subst_PTm ρ B ⟧ i ↘ S1.
 Notation "Γ ⊨ a ≲ b" := (SemLEq Γ a b) (at level 70).
 
-Lemma SemWt_Univ n Γ (A : PTm n) i  :
+Lemma SemWt_Univ Γ (A : PTm) i  :
   Γ ⊨ A ∈ PUniv i <->
   forall ρ, ρ_ok Γ ρ -> exists S, ⟦ subst_PTm ρ A ⟧ i ↘ S.
 Proof.
@@ -586,13 +585,13 @@ Proof.
     + simpl. eauto.
 Qed.
 
-Lemma SemEq_SemWt n Γ (a b A : PTm n) : Γ ⊨ a ≡ b ∈ A -> Γ ⊨ a ∈ A /\ Γ ⊨ b ∈ A /\ DJoin.R a b.
+Lemma SemEq_SemWt Γ (a b A : PTm) : Γ ⊨ a ≡ b ∈ A -> Γ ⊨ a ∈ A /\ Γ ⊨ b ∈ A /\ DJoin.R a b.
 Proof. hauto lq:on rew:off unfold:SemEq, SemWt. Qed.
 
-Lemma SemLEq_SemWt n Γ (A B : PTm n) : Γ ⊨ A ≲ B -> Sub.R A B /\ exists i, Γ ⊨ A ∈ PUniv i /\ Γ ⊨ B ∈ PUniv i.
+Lemma SemLEq_SemWt Γ (A B : PTm) : Γ ⊨ A ≲ B -> Sub.R A B /\ exists i, Γ ⊨ A ∈ PUniv i /\ Γ ⊨ B ∈ PUniv i.
 Proof. hauto q:on use:SemWt_Univ. Qed.
 
-Lemma SemWt_SemEq n Γ (a b A : PTm n) : Γ ⊨ a ∈ A -> Γ ⊨ b ∈ A -> (DJoin.R a b) -> Γ ⊨ a ≡ b ∈ A.
+Lemma SemWt_SemEq Γ (a b A : PTm) : Γ ⊨ a ∈ A -> Γ ⊨ b ∈ A -> (DJoin.R a b) -> Γ ⊨ a ≡ b ∈ A.
 Proof.
   move => ha hb heq. split => //= ρ hρ.
   have {}/ha := hρ.
@@ -603,7 +602,7 @@ Proof.
   hauto lq:on.
 Qed.
 
-Lemma SemWt_SemLEq n Γ (A B : PTm n) i j :
+Lemma SemWt_SemLEq Γ (A B : PTm) i j :
   Γ ⊨ A ∈ PUniv i -> Γ ⊨ B ∈ PUniv j -> Sub.R A B -> Γ ⊨ A ≲ B.
 Proof.
   move => ha hb heq. split => //.
@@ -621,77 +620,76 @@ Proof.
 Qed.
 
 (* Semantic context wellformedness *)
-Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ ⊨ Γ i ∈ PUniv j.
+Definition SemWff Γ := forall (i : nat) A, lookup i Γ A -> exists j, Γ ⊨ A ∈ PUniv j.
 Notation "⊨ Γ" := (SemWff Γ) (at level 70).
 
-Lemma ρ_ok_bot n (Γ : fin n -> PTm n)  :
-  ρ_ok Γ (fun _ => PBot).
+Lemma ρ_ok_id Γ :
+  ρ_ok Γ VarPTm.
 Proof.
   rewrite /ρ_ok.
   hauto q:on use:adequacy ctrs:SNe.
 Qed.
 
-Lemma ρ_ok_cons n i (Γ : fin n -> PTm n) ρ a PA A :
+Lemma ρ_ok_cons i Γ ρ a PA A :
   ⟦ subst_PTm ρ A ⟧ i ↘ PA -> PA a ->
   ρ_ok Γ ρ ->
-  ρ_ok (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ).
+  ρ_ok (cons A Γ) (scons a ρ).
 Proof.
   move => h0 h1 h2.
   rewrite /ρ_ok.
-  move => j.
-  destruct j as [j|].
+  case => [|j]; cycle 1.
   - move => m PA0. asimpl => ?.
-    asimpl.
-    firstorder.
-  - move => m PA0. asimpl => h3.
+    inversion 1; subst; asimpl.
+    hauto lq:on unfold:ρ_ok.
+  - move => m A0 PA0.
+    inversion 1; subst. asimpl => h.
     have ? : PA0 = PA by eauto using InterpUniv_Functional'.
     by subst.
 Qed.
 
-Lemma ρ_ok_cons' n i (Γ : fin n -> PTm n) ρ a PA A  Δ :
-  Δ = (funcomp (ren_PTm shift) (scons A Γ)) ->
+Lemma ρ_ok_cons' i Γ ρ a PA A  Δ :
+  Δ = (cons A Γ) ->
   ⟦ subst_PTm ρ A ⟧ i ↘ PA -> PA a ->
   ρ_ok Γ ρ ->
   ρ_ok Δ (scons a ρ).
 Proof. move => ->. apply ρ_ok_cons. Qed.
 
-Lemma ρ_ok_renaming n m (Γ : fin n -> PTm n) ρ :
-  forall (Δ : fin m -> PTm m) ξ,
+Lemma ρ_ok_renaming (Γ : list PTm) ρ :
+  forall (Δ : list PTm) ξ,
     renaming_ok Γ Δ ξ ->
     ρ_ok Γ ρ ->
     ρ_ok Δ (funcomp ρ ξ).
 Proof.
   move => Δ ξ hξ hρ.
-  rewrite /ρ_ok => i m' PA.
+  rewrite /ρ_ok => i m' A PA.
   rewrite /renaming_ok in hξ.
   rewrite /ρ_ok in hρ.
-  move => h.
+  move => PA0 h.
   rewrite /funcomp.
-  apply hρ with (k := m').
-  move : h. rewrite -hξ.
-  by asimpl.
+  eapply hρ with (k := m'); eauto.
+  move : h. by asimpl.
 Qed.
 
-Lemma renaming_SemWt {n} Γ a A :
+Lemma renaming_SemWt Γ a A :
   Γ ⊨ a ∈ A ->
-  forall {m} Δ (ξ : fin n -> fin m),
+  forall Δ (ξ : nat -> nat),
     renaming_ok Δ Γ ξ ->
     Δ ⊨ ren_PTm ξ a ∈ ren_PTm ξ A.
 Proof.
-  rewrite /SemWt => h m Δ ξ hξ ρ hρ.
+  rewrite /SemWt => h  Δ ξ hξ ρ hρ.
   have /h hρ' : (ρ_ok Γ (funcomp ρ ξ)) by eauto using ρ_ok_renaming.
   hauto q:on solve+:(by asimpl).
 Qed.
 
-Definition smorphing_ok {n m} (Δ : fin m -> PTm m) Γ (ρ : fin n -> PTm m) :=
+Definition smorphing_ok Δ Γ ρ :=
   forall ξ, ρ_ok Δ ξ -> ρ_ok Γ (funcomp (subst_PTm ξ) ρ).
 
-Lemma smorphing_ok_refl n (Δ : fin n -> PTm n) : smorphing_ok Δ Δ VarPTm.
+Lemma smorphing_ok_refl Δ : smorphing_ok Δ Δ VarPTm.
   rewrite /smorphing_ok => ξ. apply.
 Qed.
 
-Lemma smorphing_ren n m p Ξ Δ Γ
-  (ρ : fin n -> PTm m) (ξ : fin m -> fin p) :
+Lemma smorphing_ren Ξ Δ Γ
+  (ρ : nat -> PTm) (ξ : nat -> nat) :
   renaming_ok Ξ Δ ξ -> smorphing_ok Δ Γ ρ ->
   smorphing_ok Ξ Γ (funcomp (ren_PTm ξ) ρ).
 Proof.
@@ -707,111 +705,123 @@ Proof.
   move => ->. by apply hρ.
 Qed.
 
-Lemma smorphing_ext n m Δ Γ (ρ : fin n -> PTm m) (a : PTm m) (A : PTm n)  :
+Lemma smorphing_ext Δ Γ (ρ : nat -> PTm) (a : PTm) (A : PTm)  :
   smorphing_ok Δ Γ ρ ->
   Δ ⊨ a ∈ subst_PTm ρ A ->
   smorphing_ok
-  Δ (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ).
+    Δ (cons A Γ) (scons a ρ).
 Proof.
   move => h ha τ. move => /[dup] hτ.
   move : ha => /[apply].
   move => [k][PA][h0]h1.
   apply h in hτ.
-  move => i.
-  destruct i as [i|].
-  - move => k0 PA0. asimpl. rewrite {2}/funcomp. asimpl.
-    move : hτ => /[apply].
-    by asimpl.
-  - move => k0 PA0. asimpl. rewrite {2}/funcomp. asimpl.
+  case => [|i]; cycle 1.
+  - move => k0 A0 PA0. asimpl. rewrite {2}/funcomp.
+    asimpl. elim /lookup_inv => //= _.
+    move => i0 Γ0 A1 B + [?][? ?]?. subst.
+    asimpl.
+    move : hτ; by repeat move/[apply].
+  - move => k0 A0 PA0. asimpl. rewrite {2}/funcomp. asimpl.
+    elim /lookup_inv => //=_ A1 Γ0 _ [? ?] ?. subst. asimpl.
     move => *. suff : PA0 = PA by congruence.
     move : h0. asimpl.
     eauto using InterpUniv_Functional'.
 Qed.
 
-Lemma morphing_SemWt : forall n Γ (a A : PTm n),
-    Γ ⊨ a ∈ A -> forall m Δ (ρ : fin n -> PTm m),
+Lemma morphing_SemWt : forall Γ (a A : PTm ),
+    Γ ⊨ a ∈ A -> forall  Δ (ρ : nat -> PTm ),
       smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ subst_PTm ρ A.
 Proof.
-  move => n Γ a A ha m Δ ρ hρ τ hτ.
+  move => Γ a A ha Δ ρ hρ τ hτ.
   have {}/hρ {}/ha := hτ.
   asimpl. eauto.
 Qed.
 
-Lemma morphing_SemWt_Univ : forall n Γ (a : PTm n) i,
-    Γ ⊨ a ∈ PUniv i -> forall m Δ (ρ : fin n -> PTm m),
+Lemma morphing_SemWt_Univ : forall Γ (a : PTm) i,
+    Γ ⊨ a ∈ PUniv i -> forall Δ (ρ : nat -> PTm),
       smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ PUniv i.
 Proof.
-  move => n Γ a i ha.
-  move => m Δ ρ.
+  move => Γ a i ha Δ ρ.
   have -> : PUniv i = subst_PTm ρ (PUniv i) by reflexivity.
   by apply morphing_SemWt.
 Qed.
 
-Lemma weakening_Sem n Γ (a : PTm n) A B i
+Lemma weakening_Sem Γ (a : PTm) A B i
   (h0 : Γ ⊨ B ∈ PUniv i)
   (h1 : Γ ⊨ a ∈ A) :
-   funcomp (ren_PTm shift) (scons B Γ) ⊨ ren_PTm shift a ∈ ren_PTm shift A.
+   (cons B Γ) ⊨ ren_PTm shift a ∈ ren_PTm shift A.
 Proof.
   apply : renaming_SemWt; eauto.
-  hauto lq:on inv:option unfold:renaming_ok.
+  hauto lq:on ctrs:lookup unfold:renaming_ok.
 Qed.
 
-Lemma SemWt_SN n Γ (a : PTm n) A :
+Lemma weakening_Sem_Univ Γ (a : PTm) B i j
+  (h0 : Γ ⊨ B ∈ PUniv i)
+  (h1 : Γ ⊨ a ∈ PUniv j) :
+   (cons B Γ) ⊨ ren_PTm shift a ∈ PUniv j.
+Proof.
+  move : weakening_Sem h0 h1; repeat move/[apply]. done.
+Qed.
+
+Lemma SemWt_SN Γ (a : PTm) A :
   Γ ⊨ a ∈ A ->
   SN a /\ SN A.
 Proof.
   move => h.
-  have {}/h := ρ_ok_bot _ Γ => h.
-  have h0 : SN (subst_PTm (fun _ : fin n => (PBot : PTm 0)) A) by hauto l:on use:adequacy.
-  have h1 : SN (subst_PTm (fun _ : fin n => (PBot : PTm 0)) a)by hauto l:on use:adequacy.
-  move {h}. hauto l:on use:sn_unmorphing.
+  have {}/h := ρ_ok_id  Γ => h.
+  have : SN (subst_PTm VarPTm A) by hauto l:on use:adequacy.
+  have : SN (subst_PTm VarPTm a)by hauto l:on use:adequacy.
+  by asimpl.
 Qed.
 
-Lemma SemEq_SN_Join n Γ (a b A : PTm n) :
+Lemma SemEq_SN_Join Γ (a b A : PTm) :
   Γ ⊨ a ≡ b ∈ A ->
   SN a /\ SN b /\ SN A /\ DJoin.R a b.
 Proof. hauto l:on use:SemEq_SemWt, SemWt_SN. Qed.
 
-Lemma SemLEq_SN_Sub n Γ (a b : PTm n) :
+Lemma SemLEq_SN_Sub Γ (a b : PTm) :
   Γ ⊨ a ≲ b ->
   SN a /\ SN b /\ Sub.R a b.
 Proof. hauto l:on use:SemLEq_SemWt, SemWt_SN. Qed.
 
 (* Structural laws for Semantic context wellformedness *)
-Lemma SemWff_nil : SemWff null.
-Proof. case. Qed.
+Lemma SemWff_nil : SemWff nil.
+Proof. hfcrush inv:lookup. Qed.
 
-Lemma SemWff_cons n Γ (A : PTm n) i :
+Lemma SemWff_cons Γ (A : PTm) i :
     ⊨ Γ ->
     Γ ⊨ A ∈ PUniv i ->
     (* -------------- *)
-    ⊨ funcomp (ren_PTm shift) (scons A Γ).
+    ⊨ (cons A Γ).
 Proof.
   move => h h0.
-  move => j. destruct j as [j|].
-  - move /(_ j) : h => [k hk].
-    exists k. change (PUniv k) with (ren_PTm shift (PUniv k : PTm n)).
-    eauto using weakening_Sem.
+  move => j A0.  elim/lookup_inv => //=_.
   - hauto q:on use:weakening_Sem.
+  - move => i0 Γ0 A1 B + ? [*]. subst.
+    move : h => /[apply]. move => [k hk].
+    exists k. change (PUniv k) with (ren_PTm shift (PUniv k)).
+    eauto using weakening_Sem.
 Qed.
 
 (* Semantic typing rules *)
-Lemma ST_Var' n Γ (i : fin n) j :
-  Γ ⊨ Γ i ∈ PUniv j ->
-  Γ ⊨ VarPTm i ∈ Γ i.
+Lemma ST_Var' Γ (i : nat) A j :
+  lookup i Γ A ->
+  Γ ⊨ A ∈ PUniv j ->
+  Γ ⊨ VarPTm i ∈ A.
 Proof.
-  move => /SemWt_Univ h.
+  move => hl /SemWt_Univ h.
   rewrite /SemWt => ρ /[dup] hρ {}/h [S hS].
   exists j,S.
-  asimpl. firstorder.
+  asimpl. hauto q:on unfold:ρ_ok.
 Qed.
 
-Lemma ST_Var n Γ (i : fin n) :
+Lemma ST_Var Γ (i : nat) A :
   ⊨ Γ ->
-  Γ ⊨ VarPTm i ∈ Γ i.
+  lookup i Γ A ->
+  Γ ⊨ VarPTm i ∈ A.
 Proof. hauto l:on use:ST_Var' unfold:SemWff. Qed.
 
-Lemma InterpUniv_Bind_nopf n p i (A : PTm n) B PA :
+Lemma InterpUniv_Bind_nopf p i (A : PTm) B PA :
   ⟦ A ⟧ i ↘ PA ->
   (forall a, PA a -> exists PB, ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB) ->
   ⟦ PBind p A B ⟧ i ↘ (BindSpace p PA (fun a PB => ⟦ subst_PTm (scons a VarPTm) B ⟧ i ↘ PB)).
@@ -820,9 +830,9 @@ Proof.
 Qed.
 
 
-Lemma ST_Bind' n Γ i j p (A : PTm n) (B : PTm (S n)) :
+Lemma ST_Bind' Γ i j p (A : PTm) (B : PTm) :
   Γ ⊨ A ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv j ->
+  (cons A Γ) ⊨ B ∈ PUniv j ->
   Γ ⊨ PBind p A B ∈ PUniv (max i j).
 Proof.
   move => /SemWt_Univ h0 /SemWt_Univ h1.
@@ -835,9 +845,9 @@ Proof.
   - move => *. asimpl. hauto l:on use:InterpUniv_cumulative, ρ_ok_cons.
 Qed.
 
-Lemma ST_Bind n Γ i p (A : PTm n) (B : PTm (S n)) :
+Lemma ST_Bind Γ i p (A : PTm) (B : PTm) :
   Γ ⊨ A ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv i ->
+  cons A Γ ⊨ B ∈ PUniv i ->
   Γ ⊨ PBind p A B ∈ PUniv i.
 Proof.
   move => h0 h1.
@@ -846,9 +856,9 @@ Proof.
   apply ST_Bind'.
 Qed.
 
-Lemma ST_Abs n Γ (a : PTm (S n)) A B i :
+Lemma ST_Abs Γ (a : PTm) A B i :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊨ a ∈ B ->
+  (cons A Γ) ⊨ a ∈ B ->
   Γ ⊨ PAbs a ∈ PBind PPi A B.
 Proof.
   rename a into b.
@@ -868,7 +878,7 @@ Proof.
   move : ha hPA. clear. hauto q:on use:adequacy.
 Qed.
 
-Lemma ST_App n Γ (b a : PTm n) A B :
+Lemma ST_App Γ (b a : PTm) A B :
   Γ ⊨ b ∈ PBind PPi A B ->
   Γ ⊨ a ∈ A ->
   Γ ⊨ PApp b a ∈ subst_PTm (scons a VarPTm) B.
@@ -884,14 +894,14 @@ Proof.
   asimpl. hauto lq:on.
 Qed.
 
-Lemma ST_App' n Γ (b a : PTm n) A B U :
+Lemma ST_App' Γ (b a : PTm) A B U :
   U = subst_PTm (scons a VarPTm) B ->
   Γ ⊨ b ∈ PBind PPi A B ->
   Γ ⊨ a ∈ A ->
   Γ ⊨ PApp b a ∈ U.
 Proof. move => ->. apply ST_App. Qed.
 
-Lemma ST_Pair n Γ (a b : PTm n) A B i :
+Lemma ST_Pair Γ (a b : PTm) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊨ a ∈ A ->
   Γ ⊨ b ∈ subst_PTm (scons a VarPTm) B ->
@@ -917,12 +927,12 @@ Proof.
     have ? : PB0 = PB by eauto using InterpUniv_Functional'. by subst.
 Qed.
 
-Lemma N_Projs n p (a b : PTm n) :
+Lemma N_Projs p (a b : PTm) :
   rtc TRedSN a b ->
   rtc TRedSN (PProj p a) (PProj p b).
 Proof. induction 1; hauto lq:on ctrs:rtc, TRedSN. Qed.
 
-Lemma ST_Proj1 n Γ (a : PTm n) A B :
+Lemma ST_Proj1 Γ (a : PTm) A B :
   Γ ⊨ a ∈ PBind PSig A B ->
   Γ ⊨ PProj PL a ∈ A.
 Proof.
@@ -944,7 +954,7 @@ Proof.
     apply : InterpUniv_back_closs; eauto.
 Qed.
 
-Lemma ST_Proj2 n Γ (a : PTm n) A B :
+Lemma ST_Proj2 Γ (a : PTm) A B :
   Γ ⊨ a ∈ PBind PSig A B ->
   Γ ⊨ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B.
 Proof.
@@ -989,7 +999,7 @@ Proof.
     + hauto lq:on use:@relations.rtc_r, InterpUniv_back_closs.
 Qed.
 
-Lemma ST_Conv' n Γ (a : PTm n) A B i :
+Lemma ST_Conv' Γ (a : PTm) A B i :
   Γ ⊨ a ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
   Sub.R A B ->
@@ -1006,7 +1016,7 @@ Proof.
   hauto lq:on.
 Qed.
 
-Lemma ST_Conv_E n Γ (a : PTm n) A B i :
+Lemma ST_Conv_E Γ (a : PTm) A B i :
   Γ ⊨ a ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
   DJoin.R A B ->
@@ -1015,23 +1025,23 @@ Proof.
   hauto l:on use:ST_Conv', Sub.FromJoin.
 Qed.
 
-Lemma ST_Conv n Γ (a : PTm n) A B :
+Lemma ST_Conv Γ (a : PTm) A B :
   Γ ⊨ a ∈ A ->
   Γ ⊨ A ≲ B ->
   Γ ⊨ a ∈ B.
 Proof. hauto l:on use:ST_Conv', SemLEq_SemWt. Qed.
 
-Lemma SE_Refl n Γ (a : PTm n) A :
+Lemma SE_Refl Γ (a : PTm) A :
   Γ ⊨ a ∈ A ->
   Γ ⊨ a ≡ a ∈ A.
 Proof. hauto lq:on unfold:SemWt,SemEq use:DJoin.refl. Qed.
 
-Lemma SE_Symmetric n Γ (a b : PTm n) A :
+Lemma SE_Symmetric Γ (a b : PTm) A :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ b ≡ a ∈ A.
 Proof. hauto q:on unfold:SemEq. Qed.
 
-Lemma SE_Transitive n Γ (a b c : PTm n) A :
+Lemma SE_Transitive Γ (a b c : PTm) A :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ b ≡ c ∈ A ->
   Γ ⊨ a ≡ c ∈ A.
@@ -1043,36 +1053,32 @@ Proof.
   hauto l:on use:DJoin.transitive.
 Qed.
 
-Definition Γ_sub' {n} (Γ Δ : fin n -> PTm n) := forall ρ, ρ_ok Δ ρ -> ρ_ok Γ ρ.
+Definition Γ_sub' Γ Δ := forall ρ, ρ_ok Δ ρ -> ρ_ok Γ ρ.
 
-Definition Γ_eq' {n} (Γ Δ : fin n -> PTm n) := forall ρ, ρ_ok Δ ρ <-> ρ_ok Γ ρ.
+Definition Γ_eq' Γ Δ := forall ρ, ρ_ok Δ ρ <-> ρ_ok Γ ρ.
 
-Lemma Γ_sub'_refl n (Γ : fin n -> PTm n) : Γ_sub' Γ Γ.
+Lemma Γ_sub'_refl Γ : Γ_sub' Γ Γ.
   rewrite /Γ_sub'. itauto. Qed.
 
-Lemma Γ_sub'_cons n (Γ Δ : fin n -> PTm n) A B i j :
+Lemma Γ_sub'_cons Γ Δ A B i j :
   Sub.R B A ->
   Γ_sub' Γ Δ ->
   Γ ⊨ A ∈ PUniv i ->
   Δ ⊨ B ∈ PUniv j ->
-  Γ_sub' (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Δ)).
+  Γ_sub' (cons A Γ) (cons B Δ).
 Proof.
   move => hsub hsub' hA hB ρ hρ.
-
-  move => k.
-  move => k0 PA.
-  have : ρ_ok Δ (funcomp ρ shift).
+  move => k k' A0 PA.
+  have hρ_inv : ρ_ok Δ (funcomp ρ shift).
   move : hρ. clear.
   move => hρ i.
-  specialize (hρ (shift i)).
-  move => k PA. move /(_ k PA) in hρ.
-  move : hρ. asimpl. by eauto.
-  move => hρ_inv.
-  destruct k as [k|].
-  - rewrite /Γ_sub' in hsub'.
-    asimpl.
-    move /(_ (funcomp ρ shift) hρ_inv) in hsub'.
-    sfirstorder simp+:asimpl.
+  (* specialize (hρ (shift i)). *)
+  move => k A PA.
+  move /there. move /(_ B).
+  rewrite /ρ_ok in hρ.
+  move /hρ. asimpl. by apply.
+  elim /lookup_inv => //=hl.
+  move => A1 Γ0 ? [? ?] ?. subst.
   - asimpl.
     move => h.
     have {}/hsub' hρ' := hρ_inv.
@@ -1080,13 +1086,21 @@ Proof.
     move => [S]hS.
     move /SemWt_Univ : (hB) (hρ_inv)=>/[apply].
     move => [S1]hS1.
-    move /(_ None) : hρ (hS1). asimpl => /[apply].
+    move /(_ var_zero j (ren_PTm shift B) S1) : hρ (hS1). asimpl.
+    move => /[apply].
+    move /(_ ltac:(apply here)).
+    move => *.
     suff : forall x, S1 x -> PA x by firstorder.
     apply : InterpUniv_Sub; eauto.
     by apply Sub.substing.
+  - rewrite /Γ_sub' in hsub'.
+    asimpl.
+    move => i0 Γ0 A1 B0 hi0 ? [? ?]?. subst.
+    move /(_ (funcomp ρ shift) hρ_inv) in hsub'.
+    move : hsub' hi0 => /[apply]. move => /[apply]. by asimpl.
 Qed.
 
-Lemma Γ_sub'_SemWt n (Γ Δ : fin n -> PTm n) a A  :
+Lemma Γ_sub'_SemWt Γ Δ a A  :
   Γ_sub' Γ Δ ->
   Γ ⊨ a ∈ A ->
   Δ ⊨ a ∈ A.
@@ -1096,32 +1110,15 @@ Proof.
   hauto l:on.
 Qed.
 
-Definition Γ_eq {n} (Γ Δ : fin n -> PTm n)  := forall i, DJoin.R (Γ i) (Δ i).
-
-Lemma Γ_eq_ρ_ok n Γ Δ (ρ : fin n -> PTm 0) : Γ_eq Γ Δ -> ⊨ Γ -> ρ_ok Γ ρ -> ρ_ok Δ ρ.
-Proof.
-  move => hΓΔ hΓ h.
-  move => i k PA hPA.
-  move : hΓ. rewrite /SemWff. move /(_ i) => [j].
-  move => hΓ.
-  rewrite SemWt_Univ in hΓ.
-  have {}/hΓ  := h.
-  move => [S hS].
-  move /(_ i) in h. suff : PA = S by qauto l:on.
-  move : InterpUniv_Join' hPA hS. repeat move/[apply].
-  apply. move /(_ i) /DJoin.symmetric in hΓΔ.
-  hauto l:on use: DJoin.substing.
-Qed.
-
-Lemma Γ_eq_sub n (Γ Δ : fin n -> PTm n) : Γ_eq' Γ Δ <-> Γ_sub' Γ Δ /\ Γ_sub' Δ Γ.
+Lemma Γ_eq_sub Γ Δ : Γ_eq' Γ Δ <-> Γ_sub' Γ Δ /\ Γ_sub' Δ Γ.
 Proof. rewrite /Γ_eq' /Γ_sub'. hauto l:on. Qed.
 
-Lemma Γ_eq'_cons n (Γ Δ : fin n -> PTm n) A B i j :
+Lemma Γ_eq'_cons Γ Δ A B i j :
   DJoin.R B A ->
   Γ_eq' Γ Δ ->
   Γ ⊨ A ∈ PUniv i ->
   Δ ⊨ B ∈ PUniv j ->
-  Γ_eq' (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Δ)).
+  Γ_eq' (cons A Γ) (cons B Δ).
 Proof.
   move => h.
   have {h} [h0 h1] : Sub.R A B /\ Sub.R B A by hauto lq:on use:Sub.FromJoin, DJoin.symmetric.
@@ -1129,124 +1126,66 @@ Proof.
   hauto l:on use:Γ_sub'_cons.
 Qed.
 
-Lemma Γ_eq'_refl n (Γ : fin n -> PTm n) : Γ_eq' Γ Γ.
+Lemma Γ_eq'_refl Γ : Γ_eq' Γ Γ.
 Proof. rewrite /Γ_eq'. firstorder. Qed.
 
-Definition Γ_sub {n} (Γ Δ : fin n -> PTm n)  := forall i, Sub.R (Γ i) (Δ i).
-
-Lemma Γ_sub_ρ_ok n Γ Δ (ρ : fin n -> PTm 0) : Γ_sub Γ Δ -> ⊨ Γ -> ρ_ok Γ ρ -> ρ_ok Δ ρ.
-Proof.
-  move => hΓΔ hΓ h.
-  move => i k PA hPA.
-  move : hΓ. rewrite /SemWff. move /(_ i) => [j].
-  move => hΓ.
-  rewrite SemWt_Univ in hΓ.
-  have {}/hΓ  := h.
-  move => [S hS].
-  move /(_ i) in h. suff : forall x, S x -> PA x by qauto l:on.
-  move : InterpUniv_Sub hS hPA. repeat move/[apply].
-  apply. by apply Sub.substing.
-Qed.
-
-Lemma Γ_sub_refl n (Γ : fin n -> PTm n) :
-  Γ_sub Γ Γ.
-Proof. sfirstorder use:Sub.refl. Qed.
-
-Lemma Γ_sub_cons n (Γ Δ : fin n -> PTm n) A B :
-  Sub.R A B ->
-  Γ_sub Γ Δ ->
-  Γ_sub (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Δ)).
-Proof.
-  move => h h0.
-  move => i.
-  destruct i as [i|].
-  rewrite /funcomp. substify. apply Sub.substing. by asimpl.
-  rewrite /funcomp.
-  asimpl. substify. apply Sub.substing. by asimpl.
-Qed.
-
-Lemma Γ_sub_cons' n (Γ : fin n -> PTm n) A B :
-  Sub.R A B ->
-  Γ_sub (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Γ)).
-Proof. eauto using Γ_sub_refl ,Γ_sub_cons. Qed.
-
-Lemma Γ_eq_refl n (Γ : fin n -> PTm n) :
-  Γ_eq Γ Γ.
-Proof. sfirstorder use:DJoin.refl. Qed.
-
-Lemma Γ_eq_cons n (Γ Δ : fin n -> PTm n) A B :
-  DJoin.R A B ->
-  Γ_eq Γ Δ ->
-  Γ_eq (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Δ)).
-Proof.
-  move => h h0.
-  move => i.
-  destruct i as [i|].
-  rewrite /funcomp. substify. apply DJoin.substing. by asimpl.
-  rewrite /funcomp.
-  asimpl. substify. apply DJoin.substing. by asimpl.
-Qed.
-Lemma Γ_eq_cons' n (Γ : fin n -> PTm n) A B :
-  DJoin.R A B ->
-  Γ_eq (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Γ)).
-Proof. eauto using Γ_eq_refl ,Γ_eq_cons. Qed.
-
-Lemma SE_Bind' n Γ i j p (A0 A1 : PTm n) B0 B1 :
-  ⊨ Γ ->
+Lemma SE_Bind' Γ i j p (A0 A1 : PTm) B0 B1 :
   Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv j ->
+  cons A0 Γ ⊨ B0 ≡ B1 ∈ PUniv j ->
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv (max i j).
 Proof.
-  move => hΓ hA hB.
+  move => hA hB.
   apply SemEq_SemWt in hA, hB.
   apply SemWt_SemEq; last by hauto l:on use:DJoin.BindCong.
   hauto l:on use:ST_Bind'.
   apply ST_Bind'; first by tauto.
-  have hΓ' : ⊨ funcomp (ren_PTm shift) (scons A1 Γ) by hauto l:on use:SemWff_cons.
   move => ρ hρ.
-  suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
-  move : Γ_eq_ρ_ok hΓ' hρ; repeat move/[apply]. apply.
-  hauto lq:on use:Γ_eq_cons'.
+  suff : ρ_ok (cons A0 Γ) ρ by hauto l:on.
+  move : hρ.
+  suff : Γ_sub' (A0 :: Γ) (A1 :: Γ) by hauto l:on unfold:Γ_sub'.
+  apply : Γ_sub'_cons.
+  apply /Sub.FromJoin /DJoin.symmetric. tauto.
+  apply Γ_sub'_refl. hauto lq:on.
+  hauto lq:on.
 Qed.
 
-Lemma SE_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
-  ⊨ Γ ->
+Lemma SE_Bind Γ i p (A0 A1 : PTm) B0 B1 :
   Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv i ->
+  cons A0 Γ ⊨ B0 ≡ B1 ∈ PUniv i ->
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i.
 Proof.
   move => *. replace i with (max i i) by lia. auto using SE_Bind'.
 Qed.
 
-Lemma SE_Abs n Γ (a b : PTm (S n)) A B i :
+Lemma SE_Abs Γ (a b : PTm) A B i :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊨ a ≡ b ∈ B ->
+  cons A Γ ⊨ a ≡ b ∈ B ->
   Γ ⊨ PAbs a ≡ PAbs b ∈ PBind PPi A B.
 Proof.
   move => hPi /SemEq_SemWt [ha][hb]he.
   apply SemWt_SemEq; eauto using DJoin.AbsCong, ST_Abs.
 Qed.
 
-Lemma SBind_inv1 n Γ i p (A : PTm n) B :
+Lemma SBind_inv1 Γ i p (A : PTm) B :
   Γ ⊨ PBind p A B ∈ PUniv i ->
   Γ ⊨ A ∈ PUniv i.
   move /SemWt_Univ => h. apply SemWt_Univ.
   hauto lq:on rew:off  use:InterpUniv_Bind_inv.
 Qed.
 
-Lemma SE_AppEta n Γ (b : PTm n) A B i :
-  ⊨ Γ ->
+Lemma SE_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.
 Proof.
-  move => hΓ h0 h1. apply SemWt_SemEq; eauto.
+  move => h0 h1. apply SemWt_SemEq; eauto.
   apply : ST_Abs; eauto.
   have hA : Γ ⊨ A ∈ PUniv i by eauto using SBind_inv1.
-  eapply ST_App' with (A := ren_PTm shift A)(B:= ren_PTm (upRen_PTm_PTm shift) B). by asimpl.
+  eapply ST_App' with (A := ren_PTm shift A)(B:= ren_PTm (upRen_PTm_PTm shift) B). asimpl. by rewrite subst_scons_id.
   2 : {
-    apply ST_Var.
-    eauto using SemWff_cons.
+    apply : ST_Var'.
+    apply here.
+    apply : weakening_Sem_Univ; eauto.
   }
   change (PBind PPi (ren_PTm shift A) (ren_PTm (upRen_PTm_PTm shift) B)) with
     (ren_PTm shift (PBind PPi A B)).
@@ -1254,10 +1193,10 @@ Proof.
   hauto q:on ctrs:rtc,RERed.R.
 Qed.
 
-Lemma SE_AppAbs n Γ (a : PTm (S n)) b A B i:
+Lemma SE_AppAbs Γ (a : PTm) b A B i:
   Γ ⊨ PBind PPi A B ∈ PUniv i ->
   Γ ⊨ b ∈ A ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊨ a ∈ B ->
+  (cons A Γ) ⊨ a ∈ B ->
   Γ ⊨ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ subst_PTm (scons b VarPTm ) B.
 Proof.
   move => h h0 h1. apply SemWt_SemEq; eauto using ST_App, ST_Abs.
@@ -1271,7 +1210,7 @@ Proof.
   apply RRed.AppAbs.
 Qed.
 
-Lemma SE_Conv' n Γ (a b : PTm n) A B i :
+Lemma SE_Conv' Γ (a b : PTm) A B i :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
   Sub.R A B ->
@@ -1281,7 +1220,7 @@ Proof.
   apply SemWt_SemEq; eauto using ST_Conv'.
 Qed.
 
-Lemma SE_Conv n Γ (a b : PTm n) A B :
+Lemma SE_Conv Γ (a b : PTm) A B :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ A ≲ B ->
   Γ ⊨ a ≡ b ∈ B.
@@ -1290,7 +1229,7 @@ Proof.
   eauto using SE_Conv'.
 Qed.
 
-Lemma SBind_inst n Γ p i (A : PTm n) B (a : PTm n) :
+Lemma SBind_inst Γ p i (A : PTm) B (a : PTm) :
   Γ ⊨ a ∈ A ->
   Γ ⊨ PBind p A B ∈ PUniv i ->
   Γ ⊨ subst_PTm (scons a VarPTm) B ∈ PUniv i.
@@ -1310,7 +1249,7 @@ Proof.
   hauto lq:on.
 Qed.
 
-Lemma SE_Pair n Γ (a0 a1 b0 b1 : PTm n) A B i :
+Lemma SE_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 ->
@@ -1320,7 +1259,7 @@ Proof.
   apply SemWt_SemEq; eauto using ST_Pair, DJoin.PairCong, SBind_inst, DJoin.cong, ST_Conv_E, ST_Pair.
 Qed.
 
-Lemma SE_Proj1 n Γ (a b : PTm n) A B :
+Lemma SE_Proj1 Γ (a b : PTm) A B :
   Γ ⊨ a ≡ b ∈ PBind PSig A B ->
   Γ ⊨ PProj PL a ≡ PProj PL b ∈ A.
 Proof.
@@ -1328,7 +1267,7 @@ Proof.
   apply SemWt_SemEq; eauto using DJoin.ProjCong, ST_Proj1.
 Qed.
 
-Lemma SE_Proj2 n Γ i (a b : PTm n) A B :
+Lemma SE_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.
@@ -1344,22 +1283,22 @@ Proof.
 Qed.
 
 
-Lemma ST_Nat n Γ i :
-  Γ ⊨ PNat : PTm n ∈ PUniv i.
+Lemma ST_Nat Γ i :
+  Γ ⊨ PNat ∈ PUniv i.
 Proof.
   move => ?. apply SemWt_Univ. move => ρ hρ.
   eexists. by apply InterpUniv_Nat.
 Qed.
 
-Lemma ST_Zero n Γ :
-  Γ ⊨ PZero : PTm n ∈ PNat.
+Lemma ST_Zero Γ :
+  Γ ⊨ PZero ∈ PNat.
 Proof.
   move => ρ hρ. exists 0, SNat. simpl. split.
   apply InterpUniv_Nat.
   apply S_Zero.
 Qed.
 
-Lemma ST_Suc n Γ (a : PTm n) :
+Lemma ST_Suc Γ (a : PTm) :
   Γ ⊨ a ∈ PNat ->
   Γ ⊨ PSuc a ∈ PNat.
 Proof.
@@ -1372,31 +1311,31 @@ Proof.
 Qed.
 
 
-Lemma sn_unmorphing' n : (forall (a : PTm n) (s : SN a), forall m (ρ : fin m -> PTm n) b, a = subst_PTm ρ b -> SN b).
+Lemma sn_unmorphing' : (forall (a : PTm) (s : SN a), forall (ρ : nat -> PTm) b, a = subst_PTm ρ b -> SN b).
 Proof. hauto l:on use:sn_unmorphing. Qed.
 
-Lemma sn_bot_up n m (a : PTm (S n)) (ρ : fin n -> PTm m) :
-  SN (subst_PTm (scons PBot ρ) a) -> SN (subst_PTm (up_PTm_PTm ρ) a).
+Lemma sn_bot_up (a : PTm) i (ρ : nat -> PTm) :
+  SN (subst_PTm (scons (VarPTm i) ρ) a) -> SN (subst_PTm (up_PTm_PTm ρ) a).
   rewrite /up_PTm_PTm.
-  move => h. eapply sn_unmorphing' with (ρ := (scons PBot VarPTm)); eauto.
+  move => h. eapply sn_unmorphing' with (ρ := (scons (VarPTm i) VarPTm)); eauto.
   by asimpl.
 Qed.
 
-Lemma sn_bot_up2 n m (a : PTm (S (S n))) (ρ : fin n -> PTm m) :
-  SN ((subst_PTm (scons PBot (scons PBot ρ)) a)) -> SN (subst_PTm (up_PTm_PTm (up_PTm_PTm ρ)) a).
+Lemma sn_bot_up2 (a : PTm) j i (ρ : nat -> PTm) :
+  SN ((subst_PTm (scons (VarPTm j) (scons (VarPTm i) ρ)) a)) -> SN (subst_PTm (up_PTm_PTm (up_PTm_PTm ρ)) a).
   rewrite /up_PTm_PTm.
-  move => h. eapply sn_unmorphing' with (ρ := (scons PBot (scons PBot VarPTm))); eauto.
+  move => h. eapply sn_unmorphing' with (ρ := (scons (VarPTm j) (scons (VarPTm i) VarPTm))); eauto.
   by asimpl.
 Qed.
 
-Lemma SNat_SN n (a : PTm n) : SNat a -> SN a.
+Lemma SNat_SN (a : PTm) : SNat a -> SN a.
   induction 1; hauto lq:on ctrs:SN. Qed.
 
-Lemma ST_Ind s Γ P (a : PTm s) b c i :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊨ P ∈ PUniv i ->
+Lemma ST_Ind Γ P (a : PTm) b c i :
+  (cons PNat Γ) ⊨ P ∈ PUniv i ->
   Γ ⊨ a ∈ PNat ->
   Γ ⊨ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊨ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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.
 Proof.
   move => hA hc ha hb ρ hρ.
@@ -1407,18 +1346,17 @@ Proof.
   set x := PInd _ _ _ _. asimpl. subst x. move : {a} (subst_PTm ρ a) .
   move : (subst_PTm ρ b) ha1 => {}b ha1.
   move => u hu.
-  have hρb : ρ_ok (funcomp (ren_PTm shift) (scons PNat Γ)) (scons PBot ρ) by apply : ρ_ok_cons; hauto lq:on ctrs:SNat, SNe use:(@InterpUniv_Nat 0).
-
-  have hρbb : ρ_ok (funcomp (ren_PTm shift) (scons P (funcomp (ren_PTm shift) (scons PNat Γ)))) (scons PBot (scons PBot ρ)).
+  have hρb : ρ_ok (cons PNat Γ) (scons (VarPTm var_zero) ρ) by apply : ρ_ok_cons; hauto lq:on ctrs:SNat, SNe use:(@InterpUniv_Nat 0).
+  have hρbb : ρ_ok (cons P (cons PNat Γ)) (scons (VarPTm var_zero) (scons (VarPTm var_zero) ρ)).
   move /SemWt_Univ /(_ _ hρb) : hA => [S ?].
   apply : ρ_ok_cons; eauto. sauto lq:on use:adequacy.
 
   (* have snP : SN P by hauto l:on use:SemWt_SN. *)
   have snb : SN b by hauto q:on use:adequacy.
   have snP : SN (subst_PTm (up_PTm_PTm ρ) P)
-    by apply sn_bot_up; move : hA hρb => /[apply]; hauto lq:on use:adequacy.
+    by eapply sn_bot_up; move : hA hρb => /[apply]; hauto lq:on use:adequacy.
   have snc : SN (subst_PTm (up_PTm_PTm (up_PTm_PTm ρ)) c)
-    by apply sn_bot_up2; move : hb hρbb => /[apply]; hauto lq:on use:adequacy.
+    by apply: sn_bot_up2; move : hb hρbb => /[apply]; hauto lq:on use:adequacy.
 
   elim : u /hu.
   + exists m, PA. split.
@@ -1426,7 +1364,7 @@ Proof.
     * apply : InterpUniv_back_clos; eauto.
       apply N_IndZero; eauto.
   + move => a snea.
-    have hρ' : ρ_ok (funcomp (ren_PTm shift) (scons PNat Γ)) (scons a ρ)by
+    have hρ' : ρ_ok (cons PNat Γ) (scons a ρ)by
       apply : ρ_ok_cons; eauto using (InterpUniv_Nat 0); hauto ctrs:SNat.
     move /SemWt_Univ : (hA) hρ' => /[apply].
     move => [S0 hS0].
@@ -1434,7 +1372,7 @@ Proof.
     eapply adequacy; eauto.
     apply N_Ind; eauto.
   + move => a ha [j][S][h0]h1.
-    have hρ' : ρ_ok (funcomp (ren_PTm shift) (scons PNat Γ)) (scons (PSuc a) ρ)by
+    have hρ' : ρ_ok (cons PNat Γ) (scons (PSuc a) ρ)by
       apply : ρ_ok_cons; eauto using (InterpUniv_Nat 0); hauto ctrs:SNat.
     move /SemWt_Univ : (hA) (hρ') => /[apply].
     move => [S0 hS0].
@@ -1445,7 +1383,7 @@ Proof.
               (subst_PTm (up_PTm_PTm (up_PTm_PTm ρ)) c))  h1.
     move => r hr.
     have hρ'' :  ρ_ok
-                   (funcomp (ren_PTm shift) (scons P (funcomp (ren_PTm shift) (scons PNat Γ)))) (scons r (scons a ρ)) by
+                   (cons P (cons PNat Γ)) (scons r (scons a ρ)) by
       eauto using ρ_ok_cons, (InterpUniv_Nat 0).
     move : hb hρ'' => /[apply].
     move => [k][PA1][h2]h3.
@@ -1453,7 +1391,7 @@ Proof.
     have ? : PA1 = S0 by  eauto using InterpUniv_Functional'.
     by subst.
   + move => a a' hr ha' [k][PA1][h0]h1.
-    have : ρ_ok (funcomp (ren_PTm shift) (scons PNat Γ)) (scons a ρ)
+    have : ρ_ok (cons PNat Γ) (scons a ρ)
       by apply : ρ_ok_cons; hauto l:on use:S_Red,(InterpUniv_Nat 0).
     move /SemWt_Univ : hA => /[apply]. move => [PA2]h2.
     exists i, PA2. split => //.
@@ -1466,10 +1404,10 @@ Proof.
     apply RPar.morphing; last by apply RPar.refl.
     eapply LoReds.FromSN_mutual in hr.
     move /LoRed.ToRRed /RPar.FromRRed in hr.
-    hauto lq:on inv:option use:RPar.refl.
+    hauto lq:on inv:nat use:RPar.refl.
 Qed.
 
-Lemma SE_SucCong n Γ (a b : PTm n) :
+Lemma SE_SucCong Γ (a b : PTm) :
   Γ ⊨ a ≡ b ∈ PNat ->
   Γ ⊨ PSuc a ≡ PSuc b ∈ PNat.
 Proof.
@@ -1478,11 +1416,11 @@ Proof.
   hauto q:on use:REReds.suc_inv, REReds.SucCong.
 Qed.
 
-Lemma SE_IndCong n Γ P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 i :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊨ P0 ≡ P1 ∈ PUniv i ->
+Lemma SE_IndCong Γ P0 P1 (a0 a1 : PTm ) b0 b1 c0 c1 i :
+  cons PNat Γ ⊨ P0 ≡ P1 ∈ PUniv i ->
   Γ ⊨ a0 ≡ a1 ∈ PNat ->
   Γ ⊨ b0 ≡ b1 ∈ subst_PTm (scons PZero VarPTm) P0 ->
-  funcomp (ren_PTm shift) (scons P0 (funcomp (ren_PTm shift) (scons PNat Γ))) ⊨ c0 ≡ c1 ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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.
 Proof.
   move /SemEq_SemWt=>[hP0][hP1]hPe.
@@ -1501,22 +1439,22 @@ Proof.
   eapply ST_Conv_E with (i := i); eauto.
   apply : Γ_sub'_SemWt; eauto.
   apply : Γ_sub'_cons; eauto using DJoin.symmetric, Sub.FromJoin.
-  apply : Γ_sub'_cons; eauto using Sub.refl, Γ_sub'_refl, (@ST_Nat _ _ 0).
-  change (PUniv i) with (ren_PTm shift (@PUniv (S n) i)).
-  apply : weakening_Sem; eauto. move : hP1.
+  apply : Γ_sub'_cons; eauto using Sub.refl, Γ_sub'_refl, (@ST_Nat _ 0).
+  apply : weakening_Sem_Univ; eauto. move : hP1.
   move /morphing_SemWt. apply. apply smorphing_ext.
-  have -> : (funcomp VarPTm shift) = funcomp (ren_PTm shift) (@VarPTm n) by asimpl.
+  have -> : (funcomp VarPTm shift) = funcomp (ren_PTm shift) (VarPTm) by asimpl.
   apply : smorphing_ren; eauto using smorphing_ok_refl. hauto l:on inv:option.
-  apply ST_Suc. apply ST_Var' with (j := 0). apply ST_Nat.
+  apply ST_Suc. apply ST_Var' with (j := 0). apply here.
+  apply ST_Nat.
   apply DJoin.renaming. by apply DJoin.substing.
   apply : morphing_SemWt_Univ; eauto.
   apply smorphing_ext; eauto using smorphing_ok_refl.
 Qed.
 
-Lemma SE_IndZero n Γ P i (b : PTm n) c :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊨ P ∈ PUniv i ->
+Lemma SE_IndZero Γ P i (b : PTm) c :
+  (cons PNat Γ) ⊨ P ∈ PUniv i ->
   Γ ⊨ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊨ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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.
 Proof.
   move => hP hb hc.
@@ -1524,11 +1462,11 @@ Proof.
   apply DJoin.FromRRed0. apply RRed.IndZero.
 Qed.
 
-Lemma SE_IndSuc s Γ P (a : PTm s) b c i :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊨ P ∈ PUniv i ->
+Lemma SE_IndSuc Γ P (a : PTm) b c i :
+  (cons PNat Γ) ⊨ P ∈ PUniv i ->
   Γ ⊨ a ∈ PNat ->
   Γ ⊨ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊨ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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.
 Proof.
   move => hP ha hb hc.
@@ -1543,7 +1481,7 @@ Proof.
   apply RRed.IndSuc.
 Qed.
 
-Lemma SE_ProjPair1 n Γ (a b : PTm n) A B i :
+Lemma SE_ProjPair1 Γ (a b : PTm) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊨ a ∈ A ->
   Γ ⊨ b ∈ subst_PTm (scons a VarPTm) B ->
@@ -1554,7 +1492,7 @@ Proof.
   apply DJoin.FromRRed0. apply RRed.ProjPair.
 Qed.
 
-Lemma SE_ProjPair2 n Γ (a b : PTm n) A B i :
+Lemma SE_ProjPair2 Γ (a b : PTm) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊨ a ∈ A ->
   Γ ⊨ b ∈ subst_PTm (scons a VarPTm) B ->
@@ -1568,7 +1506,7 @@ Proof.
   apply DJoin.FromRRed0. apply RRed.ProjPair.
 Qed.
 
-Lemma SE_PairEta n Γ (a : PTm n) A B i :
+Lemma SE_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.
@@ -1578,7 +1516,7 @@ Proof.
   rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
 Qed.
 
-Lemma SE_Nat n Γ (a b : PTm n) :
+Lemma SE_Nat Γ (a b : PTm) :
   Γ ⊨ a ≡ b ∈ PNat ->
   Γ ⊨ PSuc a ≡ PSuc b ∈ PNat.
 Proof.
@@ -1587,7 +1525,7 @@ Proof.
   eauto using DJoin.SucCong.
 Qed.
 
-Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) A B :
+Lemma SE_App Γ i (b0 b1 a0 a1 : PTm) A B :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊨ b0 ≡ b1 ∈ PBind PPi A B ->
   Γ ⊨ a0 ≡ a1 ∈ A ->
@@ -1599,14 +1537,14 @@ Proof.
   apply : ST_Conv_E; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
 Qed.
 
-Lemma SSu_Eq n Γ (A B : PTm n) i :
+Lemma SSu_Eq Γ (A B : PTm) i :
   Γ ⊨ A ≡ B ∈ PUniv i ->
   Γ ⊨ A ≲ B.
 Proof. move /SemEq_SemWt => h.
        qauto l:on use:SemWt_SemLEq, Sub.FromJoin.
 Qed.
 
-Lemma SSu_Transitive n Γ (A B C : PTm n) :
+Lemma SSu_Transitive Γ (A B C : PTm) :
   Γ ⊨ A ≲ B ->
   Γ ⊨ B ≲ C ->
   Γ ⊨ A ≲ C.
@@ -1618,32 +1556,32 @@ Proof.
   qauto l:on use:SemWt_SemLEq, Sub.transitive.
 Qed.
 
-Lemma ST_Univ' n Γ i j :
+Lemma ST_Univ' Γ i j :
   i < j ->
-  Γ ⊨ PUniv i : PTm n ∈ PUniv j.
+  Γ ⊨ PUniv i ∈ PUniv j.
 Proof.
   move => ?.
   apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ.
 Qed.
 
-Lemma ST_Univ n Γ i :
-  Γ ⊨ PUniv i : PTm n ∈ PUniv (S i).
+Lemma ST_Univ Γ i :
+  Γ ⊨ PUniv i ∈ PUniv (S i).
 Proof.
   apply ST_Univ'. lia.
 Qed.
 
-Lemma SSu_Univ n Γ i j :
+Lemma SSu_Univ Γ i j :
   i <= j ->
-  Γ ⊨ PUniv i : PTm n ≲ PUniv j.
+  Γ ⊨ PUniv i ≲ PUniv j.
 Proof.
   move => h. apply : SemWt_SemLEq; eauto using ST_Univ.
   sauto lq:on.
 Qed.
 
-Lemma SSu_Pi n Γ (A0 A1 : PTm n) B0 B1 :
+Lemma SSu_Pi Γ (A0 A1 : PTm ) B0 B1 :
   ⊨ Γ ->
   Γ ⊨ A1 ≲ A0 ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≲ B1 ->
+  cons A0 Γ ⊨ B0 ≲ B1 ->
   Γ ⊨ PBind PPi A0 B0 ≲ PBind PPi A1 B1.
 Proof.
   move => hΓ hA hB.
@@ -1655,17 +1593,18 @@ Proof.
   apply : SemWt_SemLEq; last by hauto l:on use:Sub.PiCong.
   hauto l:on use:ST_Bind'.
   apply ST_Bind'; eauto.
-  have hΓ' : ⊨ funcomp (ren_PTm shift) (scons A1 Γ) by hauto l:on use:SemWff_cons.
+  have hΓ' : ⊨ (cons A1 Γ) by hauto l:on use:SemWff_cons.
   move => ρ hρ.
-  suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
-  move : Γ_sub_ρ_ok hΓ' hρ; repeat move/[apply]. apply.
-  hauto lq:on use:Γ_sub_cons'.
+  suff : ρ_ok (cons A0 Γ) ρ by hauto l:on.
+  move : hρ. suff : Γ_sub' (A0 :: Γ) (A1 :: Γ)
+    by hauto l:on unfold:Γ_sub'.
+  apply : Γ_sub'_cons; eauto. apply Γ_sub'_refl.
 Qed.
 
-Lemma SSu_Sig n Γ (A0 A1 : PTm n) B0 B1 :
+Lemma SSu_Sig Γ (A0 A1 : PTm) B0 B1 :
   ⊨ Γ ->
   Γ ⊨ A0 ≲ A1 ->
-  funcomp (ren_PTm shift) (scons A1 Γ) ⊨ B0 ≲ B1 ->
+  cons A1 Γ ⊨ B0 ≲ B1 ->
   Γ ⊨ PBind PSig A0 B0 ≲ PBind PSig A1 B1.
 Proof.
   move => hΓ hA hB.
@@ -1677,15 +1616,16 @@ Proof.
   apply : SemWt_SemLEq; last by hauto l:on use:Sub.SigCong.
   2 : { hauto l:on use:ST_Bind'. }
   apply ST_Bind'; eauto.
-  have hΓ' : ⊨ funcomp (ren_PTm shift) (scons A1 Γ) by hauto l:on use:SemWff_cons.
-  have hΓ'' : ⊨ funcomp (ren_PTm shift) (scons A0 Γ) by hauto l:on use:SemWff_cons.
+  have hΓ' : ⊨ cons A1 Γ by hauto l:on use:SemWff_cons.
+  have hΓ'' : ⊨ cons A0 Γ by hauto l:on use:SemWff_cons.
   move => ρ hρ.
-  suff : ρ_ok (funcomp (ren_PTm shift) (scons A1 Γ)) ρ by hauto l:on.
-  apply : Γ_sub_ρ_ok; eauto.
-  hauto lq:on use:Γ_sub_cons'.
+  suff : ρ_ok (cons A1 Γ) ρ by hauto l:on.
+  move : hρ. suff : Γ_sub' (A1 :: Γ) (A0 :: Γ) by hauto l:on.
+  apply : Γ_sub'_cons; eauto.
+  apply Γ_sub'_refl.
 Qed.
 
-Lemma SSu_Pi_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+Lemma SSu_Pi_Proj1 Γ (A0 A1 : PTm) B0 B1 :
   Γ ⊨ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
   Γ ⊨ A1 ≲ A0.
 Proof.
@@ -1694,7 +1634,7 @@ Proof.
   hauto lq:on rew:off use:Sub.bind_inj.
 Qed.
 
-Lemma SSu_Sig_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+Lemma SSu_Sig_Proj1 Γ (A0 A1 : PTm) B0 B1 :
   Γ ⊨ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
   Γ ⊨ A0 ≲ A1.
 Proof.
@@ -1703,7 +1643,7 @@ Proof.
   hauto lq:on rew:off use:Sub.bind_inj.
 Qed.
 
-Lemma SSu_Pi_Proj2 n Γ (a0 a1 A0 A1 : PTm n) B0 B1 :
+Lemma SSu_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.
@@ -1716,7 +1656,7 @@ Proof.
   apply : ST_Conv'; eauto. hauto l:on use:SBind_inv1.
 Qed.
 
-Lemma SSu_Sig_Proj2 n Γ (a0 a1 A0 A1 : PTm n) B0 B1 :
+Lemma SSu_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.

From b3bd75ad429cb400cb5a8ab38446dd63dc8c9fb9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 3 Mar 2025 01:38:22 -0500
Subject: [PATCH 06/24] Fix the typing rules

---
 theories/common.v     |   7 ++
 theories/structural.v |  77 ++++++++--------------
 theories/typing.v     | 145 +++++++++++++++++++++---------------------
 3 files changed, 103 insertions(+), 126 deletions(-)

diff --git a/theories/common.v b/theories/common.v
index a890edd..a3740ff 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -10,6 +10,13 @@ Inductive lookup : nat -> list PTm -> PTm -> Prop :=
   lookup i Γ A ->
   lookup (S i) (cons B Γ) (ren_PTm shift A).
 
+Lemma here' A Γ U : U = ren_PTm shift A -> lookup 0 (A :: Γ) U.
+Proof.  move => ->. apply here. Qed.
+
+Lemma there' i Γ A B U : U = ren_PTm shift A -> lookup i Γ A ->
+                       lookup (S i) (cons B Γ) U.
+Proof. move => ->. apply there. Qed.
+
 Derive Inversion lookup_inv with (forall i Γ A, lookup i Γ A).
 
 Definition renaming_ok (Γ : list PTm) (Δ : list PTm) (ξ : nat -> nat) :=
diff --git a/theories/structural.v b/theories/structural.v
index c25986c..39a573d 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -1,96 +1,69 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common typing.
 From Hammer Require Import Tactics.
 Require Import ssreflect.
 Require Import Psatz.
 
 Lemma wff_mutual :
-  (forall n (Γ : fin n -> PTm n), ⊢ Γ -> True) /\
-  (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 -> ⊢ Γ).
+  (forall Γ, ⊢ Γ -> True) /\
+  (forall Γ (a A : PTm), Γ ⊢ a ∈ A -> ⊢ Γ)  /\
+  (forall Γ (a b A : PTm), Γ ⊢ a ≡ b ∈ A -> ⊢ Γ) /\
+  (forall Γ (A B : PTm), Γ ⊢ A ≲ B -> ⊢ Γ).
 Proof. apply wt_mutual; eauto. Qed.
 
 #[export]Hint Constructors Wt Wff Eq : wt.
 
-Lemma T_Nat' n Γ :
+Lemma T_Nat' Γ :
   ⊢ Γ ->
-  Γ ⊢ PNat : PTm n ∈ PUniv 0.
+  Γ ⊢ PNat ∈ PUniv 0.
 Proof. apply T_Nat. Qed.
 
-Lemma renaming_up n m (ξ : fin n -> fin m) Δ Γ A :
+Lemma renaming_up (ξ : nat -> nat) Δ Γ A :
   renaming_ok Δ Γ ξ ->
-  renaming_ok (funcomp (ren_PTm shift) (scons (ren_PTm ξ A) Δ)) (funcomp (ren_PTm shift) (scons A Γ)) (upRen_PTm_PTm ξ) .
+  renaming_ok (cons (ren_PTm ξ A) Δ) (cons A Γ) (upRen_PTm_PTm ξ) .
 Proof.
-  move => h i.
-  destruct i as [i|].
-  asimpl. rewrite /renaming_ok in h.
-  rewrite /funcomp. rewrite -h.
-  by asimpl.
-  by asimpl.
+  move => h i A0.
+  elim /lookup_inv => //=_.
+  - move => A1 Γ0 ? [*]. subst. apply here'. by asimpl.
+  - move => i0 Γ0 A1 B h' ? [*]. subst.
+    apply : there'; eauto. by asimpl.
 Qed.
 
-Lemma Su_Wt n Γ a i :
-  Γ ⊢ a ∈ @PUniv n i ->
+Lemma Su_Wt Γ a i :
+  Γ ⊢ a ∈ PUniv i ->
   Γ ⊢ a ≲ a.
 Proof. hauto lq:on ctrs:LEq, Eq. Qed.
 
-Lemma Wt_Univ n Γ a A i
+Lemma Wt_Univ Γ a A i
   (h : Γ ⊢ a ∈ A) :
-  Γ ⊢ @PUniv n i ∈ PUniv (S i).
+  Γ ⊢ @PUniv i ∈ PUniv (S i).
 Proof.
   hauto lq:on ctrs:Wt use:wff_mutual.
 Qed.
 
-Lemma Bind_Inv n Γ p (A : PTm n) B U :
+Lemma Bind_Inv Γ p (A : PTm) B U :
   Γ ⊢ PBind p A B ∈ U ->
   exists i, Γ ⊢ A ∈ PUniv i /\
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i /\
+  (cons A Γ) ⊢ B ∈ PUniv i /\
   Γ ⊢ PUniv i ≲ U.
 Proof.
   move E :(PBind p A B) => T h.
   move : p A B E.
-  elim : n Γ T U / h => //=.
-  - move => n Γ i p A B hA _ hB _ p0 A0 B0 [*]. subst.
+  elim : Γ T U / h => //=.
+  - move => Γ i p A B hA _ hB _ p0 A0 B0 [*]. subst.
     exists i. repeat split => //=.
     eapply wff_mutual in hA.
     apply Su_Univ; eauto.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-(* Lemma Pi_Inv n Γ (A : PTm n) B U : *)
-(*   Γ ⊢ PBind PPi A B ∈ U -> *)
-(*   exists i, Γ ⊢ A ∈ PUniv i /\ *)
-(*   funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i /\ *)
-(*   Γ ⊢ PUniv i ≲ U. *)
-(* Proof. *)
-(*   move E :(PBind PPi A B) => T h. *)
-(*   move : A B E. *)
-(*   elim : n Γ T U / h => //=. *)
-(*   - hauto lq:on ctrs:Wt,LEq,Eq use:Wt_Univ. *)
-(*   - hauto lq:on rew:off ctrs:LEq. *)
-(* Qed. *)
-
-(* Lemma Bind_Inv n Γ (A : PTm n) B U : *)
-(*   Γ ⊢ PBind PSig A B ∈ U -> *)
-(*   exists i, Γ ⊢ A ∈ PUniv i /\ *)
-(*   funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i /\ *)
-(*   Γ ⊢ PUniv i ≲ U. *)
-(* Proof. *)
-(*   move E :(PBind PSig A B) => T h. *)
-(*   move : A B E. *)
-(*   elim : n Γ T U / h => //=. *)
-(*   - hauto lq:on ctrs:Wt,LEq,Eq use:Wt_Univ. *)
-(*   - hauto lq:on rew:off ctrs:LEq. *)
-(* Qed. *)
-
-Lemma T_App' n Γ (b a : PTm n) A B U :
+Lemma T_App' Γ (b a : PTm) A B U :
   U = subst_PTm (scons a VarPTm) B ->
   Γ ⊢ b ∈ PBind PPi A B ->
   Γ ⊢ a ∈ A ->
   Γ ⊢ PApp b a ∈ U.
 Proof. move => ->. apply T_App. Qed.
 
-Lemma T_Pair' n Γ (a b : PTm n) A B i U :
+Lemma T_Pair' Γ (a b : PTm ) A B i U :
   U = subst_PTm (scons a VarPTm) B ->
   Γ ⊢ a ∈ A ->
   Γ ⊢ b ∈ U ->
@@ -100,7 +73,7 @@ Proof.
   move => ->. eauto using T_Pair.
 Qed.
 
-Lemma E_IndCong' n Γ P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 i U :
+Lemma E_IndCong' Γ P0 P1 (a0 a1 : PTm ) b0 b1 c0 c1 i U :
   U = subst_PTm (scons a0 VarPTm) P0 ->
   funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P0 ∈ PUniv i ->
   funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P0 ≡ P1 ∈ PUniv i ->
diff --git a/theories/typing.v b/theories/typing.v
index 818d6b5..ae78747 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -1,240 +1,237 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 
 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 : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
-| T_Var n Γ (i : fin n) :
+Inductive Wt : list PTm -> PTm -> PTm -> Prop :=
+| T_Var i Γ A :
   ⊢ Γ ->
-  Γ ⊢ VarPTm i ∈ Γ i
+  lookup i Γ A ->
+  Γ ⊢ VarPTm i ∈ A
 
-| T_Bind n Γ i p (A : PTm n) (B : PTm (S n)) :
+| T_Bind Γ i p (A : PTm) (B : PTm) :
   Γ ⊢ A ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i ->
+  cons A Γ ⊢ B ∈ PUniv i ->
   Γ ⊢ PBind p A B ∈ PUniv i
 
-| T_Abs n Γ (a : PTm (S n)) A B i :
+| T_Abs Γ (a : PTm) A B i :
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
+  (cons A Γ) ⊢ a ∈ B ->
   Γ ⊢ PAbs a ∈ PBind PPi A B
 
-| T_App n Γ (b a : PTm n) 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 n Γ (a b : PTm n) A B i :
+| 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 n Γ (a : PTm n) A B :
+| T_Proj1 Γ (a : PTm) A B :
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PL a ∈ A
 
-| T_Proj2 n Γ (a : PTm n) A B :
+| T_Proj2 Γ (a : PTm) A B :
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B
 
-| T_Univ n Γ i :
+| T_Univ Γ i :
   ⊢ Γ ->
-  Γ ⊢ PUniv i : PTm n ∈ PUniv (S i)
+  Γ ⊢ PUniv i ∈ PUniv (S i)
 
-| T_Nat n Γ i :
+| T_Nat Γ i :
   ⊢ Γ ->
-  Γ ⊢ PNat : PTm n ∈ PUniv i
+  Γ ⊢ PNat ∈ PUniv i
 
-| T_Zero n Γ :
+| T_Zero Γ :
   ⊢ Γ ->
-  Γ ⊢ PZero : PTm n ∈ PNat
+  Γ ⊢ PZero ∈ PNat
 
-| T_Suc n Γ (a : PTm n) :
+| T_Suc Γ (a : PTm) :
   Γ ⊢ a ∈ PNat ->
   Γ ⊢ PSuc a ∈ PNat
 
-| T_Ind s Γ P (a : PTm s) b c i :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
+| T_Ind Γ P (a : PTm) b c i :
+  cons PNat Γ ⊢ P ∈ PUniv i ->
   Γ ⊢ a ∈ PNat ->
   Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 n Γ (a : PTm n) A B :
+| T_Conv Γ (a : PTm) A B :
   Γ ⊢ a ∈ A ->
   Γ ⊢ A ≲ B ->
   Γ ⊢ a ∈ B
 
-with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
+with Eq : list PTm -> PTm -> PTm -> PTm -> Prop :=
 (* Structural *)
-| E_Refl n Γ (a : PTm n) A :
+| E_Refl Γ (a : PTm ) A :
   Γ ⊢ a ∈ A ->
   Γ ⊢ a ≡ a ∈ A
 
-| E_Symmetric n Γ (a b : PTm n) A :
+| E_Symmetric Γ (a b : PTm) A :
   Γ ⊢ a ≡ b ∈ A ->
   Γ ⊢ b ≡ a ∈ A
 
-| E_Transitive n Γ (a b c : PTm n) A :
+| E_Transitive Γ (a b c : PTm) A :
   Γ ⊢ a ≡ b ∈ A ->
   Γ ⊢ b ≡ c ∈ A ->
   Γ ⊢ a ≡ c ∈ A
 
 (* Congruence *)
-| E_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
-  ⊢ Γ ->
+| E_Bind Γ i p (A0 A1 : PTm) B0 B1 :
   Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
+  (cons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i
 
-| E_Abs n Γ (a b : PTm (S n)) A B i :
+| E_Abs Γ (a b : PTm) A B i :
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ≡ b ∈ B ->
+  (cons A Γ) ⊢ a ≡ b ∈ B ->
   Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B
 
-| E_App n Γ i (b0 b1 a0 a1 : PTm n) 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 n Γ (a0 a1 b0 b1 : PTm n) A B i :
+| 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 n Γ (a b : PTm n) A B :
+| E_Proj1 Γ (a b : PTm) A B :
   Γ ⊢ a ≡ b ∈ PBind PSig A B ->
   Γ ⊢ PProj PL a ≡ PProj PL b ∈ A
 
-| E_Proj2 n Γ i (a b : PTm n) A B :
+| 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 n Γ P0 P1 (a0 a1 : PTm n) b0 b1 c0 c1 i :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P0 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P0 ≡ P1 ∈ PUniv i ->
+| 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 ->
-  funcomp (ren_PTm shift) (scons P0 (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c0 ≡ c1 ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 n Γ (a b : PTm n) :
+| E_SucCong Γ (a b : PTm) :
   Γ ⊢ a ≡ b ∈ PNat ->
   Γ ⊢ PSuc a ≡ PSuc b ∈ PNat
 
-| E_Conv n Γ (a b : PTm n) A B :
+| E_Conv Γ (a b : PTm) A B :
   Γ ⊢ a ≡ b ∈ A ->
   Γ ⊢ A ≲ B ->
   Γ ⊢ a ≡ b ∈ B
 
 (* Beta *)
-| E_AppAbs n Γ (a : PTm (S n)) b A B i:
+| E_AppAbs Γ (a : PTm) b A B i:
   Γ ⊢ PBind PPi A B ∈ PUniv i ->
   Γ ⊢ b ∈ A ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
+  (cons A Γ) ⊢ a ∈ B ->
   Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ subst_PTm (scons b VarPTm ) B
 
-| E_ProjPair1 n Γ (a b : PTm n) A B i :
+| 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 n Γ (a b : PTm n) A B i :
+| 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 n Γ P i (b : PTm n) c :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
+| E_IndZero Γ P i (b : PTm) c :
+  (cons PNat Γ) ⊢ P ∈ PUniv i ->
   Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 s Γ P (a : PTm s) b c i :
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
+| E_IndSuc Γ P (a : PTm) b c i :
+  (cons PNat Γ) ⊢ P ∈ PUniv i ->
   Γ ⊢ a ∈ PNat ->
   Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 n Γ (b : PTm n) A B i :
-  ⊢ Γ ->
+| 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 n Γ (a : PTm n) A B i :
+| 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 : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
+with LEq : list PTm -> PTm -> PTm -> Prop :=
 (* Structural *)
-| Su_Transitive  n Γ (A B C : PTm n) :
+| Su_Transitive Γ (A B C : PTm) :
   Γ ⊢ A ≲ B ->
   Γ ⊢ B ≲ C ->
   Γ ⊢ A ≲ C
 
 (* Congruence *)
-| Su_Univ n Γ i j :
+| Su_Univ Γ i j :
   ⊢ Γ ->
   i <= j ->
-  Γ ⊢ PUniv i : PTm n ≲ PUniv j
+  Γ ⊢ PUniv i ≲ PUniv j
 
-| Su_Pi n Γ (A0 A1 : PTm n) B0 B1 i :
-  ⊢ Γ ->
+| Su_Pi Γ (A0 A1 : PTm) B0 B1 i :
   Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A1 ≲ A0 ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≲ B1 ->
+  (cons A0 Γ) ⊢ B0 ≲ B1 ->
   Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1
 
-| Su_Sig n Γ (A0 A1 : PTm n) B0 B1 i :
-  ⊢ Γ ->
+| Su_Sig Γ (A0 A1 : PTm) B0 B1 i :
   Γ ⊢ A1 ∈ PUniv i ->
   Γ ⊢ A0 ≲ A1 ->
-  funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ≲ B1 ->
+  (cons A1 Γ) ⊢ B0 ≲ B1 ->
   Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1
 
 (* Injecting from equalities *)
-| Su_Eq n Γ (A : PTm n) B i  :
+| Su_Eq Γ (A : PTm) B i  :
   Γ ⊢ A ≡ B ∈ PUniv i ->
   Γ ⊢ A ≲ B
 
 (* Projection axioms *)
-| Su_Pi_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+| Su_Pi_Proj1 Γ (A0 A1 : PTm) B0 B1 :
   Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
   Γ ⊢ A1 ≲ A0
 
-| Su_Sig_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+| Su_Sig_Proj1 Γ (A0 A1 : PTm) B0 B1 :
   Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
   Γ ⊢ A0 ≲ A1
 
-| Su_Pi_Proj2 n Γ (a0 a1 A0 A1 : PTm n) B0 B1 :
+| 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 n Γ (a0 a1 A0 A1 : PTm n) B0 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 : forall {n}, (fin n -> PTm n) -> Prop :=
+with Wff : list PTm -> Prop :=
 | Wff_Nil :
-  ⊢ null
-| Wff_Cons n Γ (A : PTm n) i :
+  ⊢ nil
+| Wff_Cons Γ (A : PTm) i :
   ⊢ Γ ->
   Γ ⊢ A ∈ PUniv i ->
   (* -------------------------------- *)
-  ⊢ funcomp (ren_PTm shift) (scons A Γ)
+  ⊢ (cons A Γ)
 
 where
 "Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A) and "Γ ⊢ A ≲ B" := (LEq Γ A B).

From 896d22ac9bd010ca83332c8b6d876ac4d314a7a5 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 3 Mar 2025 01:40:12 -0500
Subject: [PATCH 07/24] minor

---
 theories/logrel.v | 4 +---
 1 file changed, 1 insertion(+), 3 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index a113437..ae4169c 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1579,7 +1579,6 @@ Proof.
 Qed.
 
 Lemma SSu_Pi Γ (A0 A1 : PTm ) B0 B1 :
-  ⊨ Γ ->
   Γ ⊨ A1 ≲ A0 ->
   cons A0 Γ ⊨ B0 ≲ B1 ->
   Γ ⊨ PBind PPi A0 B0 ≲ PBind PPi A1 B1.
@@ -1602,12 +1601,11 @@ Proof.
 Qed.
 
 Lemma SSu_Sig Γ (A0 A1 : PTm) B0 B1 :
-  ⊨ Γ ->
   Γ ⊨ A0 ≲ A1 ->
   cons A1 Γ ⊨ B0 ≲ B1 ->
   Γ ⊨ PBind PSig A0 B0 ≲ PBind PSig A1 B1.
 Proof.
-  move => hΓ hA hB.
+  move => hA hB.
   have ? : SN A0 /\ SN A1 /\ SN B0 /\ SN B1
     by hauto l:on use:SemLEq_SN_Sub.
   apply SemLEq_SemWt in hA, hB.

From d68adf85f4b9a39ab24e5d2b814bf485b3fcad27 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 15:22:59 -0500
Subject: [PATCH 08/24] Finish refactoring substitution lemmas

---
 theories/common.v       |  16 +-
 theories/fp_red.v       |   8 -
 theories/preservation.v |  12 +-
 theories/structural.v   | 443 ++++++++++++++++++++--------------------
 4 files changed, 244 insertions(+), 235 deletions(-)

diff --git a/theories/common.v b/theories/common.v
index a3740ff..5095fef 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -10,6 +10,12 @@ Inductive lookup : nat -> list PTm -> PTm -> Prop :=
   lookup i Γ A ->
   lookup (S i) (cons B Γ) (ren_PTm shift A).
 
+Lemma lookup_deter i Γ A B :
+  lookup i Γ A ->
+  lookup i Γ B ->
+  A = B.
+Proof. move => h. move : B. induction h; hauto lq:on inv:lookup. Qed.
+
 Lemma here' A Γ U : U = ren_PTm shift A -> lookup 0 (A :: Γ) U.
 Proof.  move => ->. apply here. Qed.
 
@@ -126,6 +132,14 @@ Definition ishne_ren (a : PTm)  (ξ : nat -> nat) :
   ishne (ren_PTm ξ a) = ishne a.
 Proof. move : ξ. elim : a => //=. Qed.
 
-Lemma renaming_shift Γ (ρ : nat -> PTm) A :
+Lemma renaming_shift Γ A :
   renaming_ok (cons A Γ) Γ shift.
 Proof. rewrite /renaming_ok. hauto lq:on ctrs:lookup. Qed.
+
+Lemma subst_scons_id (a : PTm) :
+  subst_PTm (scons (VarPTm 0) (funcomp VarPTm shift)) a = a.
+Proof.
+  have E : subst_PTm VarPTm a = a by asimpl.
+  rewrite -{2}E.
+  apply ext_PTm. case => //=.
+Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 6aa7cb6..9d8718b 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -12,14 +12,6 @@ Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 Require Import Btauto.
 Require Import Cdcl.Itauto.
 
-Lemma subst_scons_id (a : PTm) :
-  subst_PTm (scons (VarPTm 0) (funcomp VarPTm shift)) a = a.
-Proof.
-  have E : subst_PTm VarPTm a = a by asimpl.
-  rewrite -{2}E.
-  apply ext_PTm. case => //=.
-Qed.
-
 Ltac2 spec_refl () :=
   List.iter
     (fun a => match a with
diff --git a/theories/preservation.v b/theories/preservation.v
index b78f87e..fe0a920 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -1,15 +1,15 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing structural fp_red.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common typing structural fp_red.
 From Hammer Require Import Tactics.
 Require Import ssreflect.
 Require Import Psatz.
 Require Import Coq.Logic.FunctionalExtensionality.
 
-Lemma App_Inv n Γ (b a : PTm n) U :
+Lemma App_Inv Γ (b a : PTm) U :
   Γ ⊢ PApp b a ∈ U ->
   exists A B, Γ ⊢ b ∈ PBind PPi A B /\ Γ ⊢ a ∈ A /\ Γ ⊢ subst_PTm (scons a VarPTm) B ≲ U.
 Proof.
   move E : (PApp b a) => u hu.
-  move : b a E. elim : n Γ u U / hu => n //=.
+  move : b a E. elim : Γ u U / hu => //=.
   - move => Γ b a A B hb _ ha _ b0 a0 [*]. subst.
     exists A,B.
     repeat split => //=.
@@ -18,12 +18,12 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma Abs_Inv n Γ (a : PTm (S n)) U :
+Lemma Abs_Inv Γ (a : PTm) U :
   Γ ⊢ PAbs a ∈ U ->
-  exists A B, funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B /\ Γ ⊢ PBind PPi A B ≲ U.
+  exists A B, (cons A Γ) ⊢ a ∈ B /\ Γ ⊢ PBind PPi A B ≲ U.
 Proof.
   move E : (PAbs a) => u hu.
-  move : a E. elim : n Γ u U / hu => n //=.
+  move : a E. elim : Γ u U / hu => //=.
   - move => Γ a A B i hP _ ha _ a0 [*]. subst.
     exists A, B. repeat split => //=.
     hauto lq:on use:E_Refl, Su_Eq.
diff --git a/theories/structural.v b/theories/structural.v
index 39a573d..10d6583 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -75,44 +75,44 @@ Qed.
 
 Lemma E_IndCong' Γ P0 P1 (a0 a1 : PTm ) b0 b1 c0 c1 i U :
   U = subst_PTm (scons a0 VarPTm) P0 ->
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P0 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P0 ≡ P1 ∈ PUniv i ->
+  (cons PNat Γ) ⊢ P0 ∈ PUniv i ->
+  (cons PNat Γ) ⊢ P0 ≡ P1 ∈ PUniv i ->
   Γ ⊢ a0 ≡ a1 ∈ PNat ->
   Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons PZero VarPTm) P0 ->
-  funcomp (ren_PTm shift) (scons P0 (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c0 ≡ c1 ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 ∈ U.
 Proof. move => ->. apply E_IndCong. Qed.
 
-Lemma T_Ind' s Γ P (a : PTm s) b c i U :
+Lemma T_Ind' Γ P (a : PTm) b c i U :
   U = subst_PTm (scons a VarPTm) P ->
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
+  cons PNat Γ ⊢ P ∈ PUniv i ->
   Γ ⊢ a ∈ PNat ->
   Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 ∈ U.
 Proof. move =>->. apply T_Ind. Qed.
 
-Lemma T_Proj2' n Γ (a : PTm n) A B U :
+Lemma T_Proj2' Γ (a : PTm) A B U :
   U = subst_PTm (scons (PProj PL a) VarPTm) B ->
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ∈ U.
 Proof. move => ->. apply T_Proj2. Qed.
 
-Lemma E_Proj2' n Γ i (a b : PTm n) A B U :
+Lemma E_Proj2' Γ i (a b : PTm) A B U :
   U = subst_PTm (scons (PProj PL a) VarPTm) B ->
   Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊢ a ≡ b ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ≡ PProj PR b ∈ U.
 Proof. move => ->. apply E_Proj2. Qed.
 
-Lemma E_Bind' n Γ i p (A0 A1 : PTm n) B0 B1 :
+Lemma E_Bind' Γ i p (A0 A1 : PTm) B0 B1 :
   Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
+  cons A0 Γ ⊢ B0 ≡ B1 ∈ PUniv i ->
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i.
 Proof. hauto lq:on use:E_Bind, wff_mutual. Qed.
 
-Lemma E_App' n Γ i (b0 b1 a0 a1 : PTm n) A B U :
+Lemma E_App' Γ i (b0 b1 a0 a1 : PTm) A B U :
   U = subst_PTm (scons a0 VarPTm) B ->
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B ->
@@ -120,16 +120,16 @@ Lemma E_App' n Γ i (b0 b1 a0 a1 : PTm n) A B U :
   Γ ⊢ PApp b0 a0 ≡ PApp b1 a1 ∈ U.
 Proof. move => ->. apply E_App. Qed.
 
-Lemma E_AppAbs' n Γ (a : PTm (S n)) b A B i u U :
+Lemma E_AppAbs' Γ (a : PTm) b A B i u U :
   u = subst_PTm (scons b VarPTm) a ->
   U = subst_PTm (scons b VarPTm ) B ->
   Γ ⊢ PBind PPi A B ∈ PUniv i ->
   Γ ⊢ b ∈ A ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
+  cons A Γ ⊢ a ∈ B ->
   Γ ⊢ PApp (PAbs a) b ≡ u ∈ U.
   move => -> ->. apply E_AppAbs. Qed.
 
-Lemma E_ProjPair2' n Γ (a b : PTm n) A B i U :
+Lemma E_ProjPair2' Γ (a b : PTm) A B i U :
   U = subst_PTm (scons a VarPTm) B ->
   Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊢ a ∈ A ->
@@ -137,14 +137,14 @@ Lemma E_ProjPair2' n Γ (a b : PTm n) A B i U :
   Γ ⊢ PProj PR (PPair a b) ≡ b ∈ U.
 Proof. move => ->. apply E_ProjPair2. Qed.
 
-Lemma E_AppEta' n Γ (b : PTm n) A B i u :
+Lemma E_AppEta' Γ (b : PTm ) A B i u :
   u = (PApp (ren_PTm shift b) (VarPTm var_zero)) ->
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊢ b ∈ PBind PPi A B ->
   Γ ⊢ PAbs u ≡ b ∈ PBind PPi A B.
 Proof. qauto l:on use:wff_mutual, E_AppEta. Qed.
 
-Lemma Su_Pi_Proj2' n Γ (a0 a1 A0 A1 : PTm n) B0 B1 U T :
+Lemma Su_Pi_Proj2' Γ (a0 a1 A0 A1 : PTm ) B0 B1 U T :
   U = subst_PTm (scons a0 VarPTm) B0 ->
   T = subst_PTm (scons a1 VarPTm) B1 ->
   Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
@@ -152,7 +152,7 @@ Lemma Su_Pi_Proj2' n Γ (a0 a1 A0 A1 : PTm n) B0 B1 U T :
   Γ ⊢ U ≲ T.
 Proof. move => -> ->. apply Su_Pi_Proj2. Qed.
 
-Lemma Su_Sig_Proj2' n Γ (a0 a1 A0 A1 : PTm n) B0 B1 U T :
+Lemma Su_Sig_Proj2' Γ (a0 a1 A0 A1 : PTm) B0 B1 U T :
   U = subst_PTm (scons a0 VarPTm) B0 ->
   T = subst_PTm (scons a1 VarPTm) B1 ->
   Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
@@ -160,64 +160,61 @@ Lemma Su_Sig_Proj2' n Γ (a0 a1 A0 A1 : PTm n) B0 B1 U T :
   Γ ⊢ U ≲ T.
 Proof. move => -> ->. apply Su_Sig_Proj2. Qed.
 
-Lemma E_IndZero' n Γ P i (b : PTm n) c U :
+Lemma E_IndZero' Γ P i (b : PTm) c U :
   U = subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
+  cons PNat Γ ⊢ P ∈ PUniv i ->
   Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 ∈ U.
 Proof. move => ->. apply E_IndZero. Qed.
 
-Lemma Wff_Cons' n Γ (A : PTm n) i :
+Lemma Wff_Cons' Γ (A : PTm) i :
   Γ ⊢ A ∈ PUniv i ->
   (* -------------------------------- *)
-  ⊢ funcomp (ren_PTm shift) (scons A Γ).
+  ⊢ cons A Γ.
 Proof. hauto lq:on rew:off use:Wff_Cons, wff_mutual. Qed.
 
-Lemma E_IndSuc' s Γ P (a : PTm s) b c i u U :
+Lemma E_IndSuc' Γ P (a : PTm) b c i u U :
   u = subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c ->
   U = subst_PTm (scons (PSuc a) VarPTm) P ->
-  funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
+  cons PNat Γ ⊢ P ∈ PUniv i ->
   Γ ⊢ a ∈ PNat ->
   Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P ->
-  funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) 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 ≡ u ∈ U.
 Proof. move => ->->. apply E_IndSuc. Qed.
 
 Lemma renaming :
-  (forall n (Γ : fin n -> PTm n), ⊢ Γ -> True) /\
-  (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> forall m Δ (ξ : fin n -> fin m), ⊢ Δ -> renaming_ok Δ Γ ξ ->
+  (forall Γ, ⊢ Γ -> True) /\
+  (forall Γ (a A : PTm), Γ ⊢ a ∈ A -> forall Δ (ξ : nat -> nat), ⊢ Δ -> renaming_ok Δ Γ ξ ->
      Δ ⊢ ren_PTm ξ a ∈ ren_PTm ξ A) /\
-  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> forall m Δ (ξ : fin n -> fin m), ⊢ Δ -> renaming_ok Δ Γ ξ ->
+  (forall Γ (a b A : PTm), Γ ⊢ a ≡ b ∈ A -> forall Δ (ξ : nat -> nat), ⊢ Δ -> renaming_ok Δ Γ ξ ->
      Δ ⊢ ren_PTm ξ a ≡ ren_PTm ξ b ∈ ren_PTm ξ A) /\
-  (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> forall  m Δ (ξ : fin n -> fin m), ⊢ Δ -> renaming_ok Δ Γ ξ ->
+  (forall Γ (A B : PTm), Γ ⊢ A ≲ B -> forall  Δ (ξ : nat -> nat), ⊢ Δ -> renaming_ok Δ Γ ξ ->
     Δ ⊢ ren_PTm ξ A ≲ ren_PTm ξ B).
 Proof.
   apply wt_mutual => //=; eauto 3 with wt.
-  - move => n Γ i hΓ _ m Δ ξ hΔ hξ.
-    rewrite hξ.
-    by apply T_Var.
   - hauto lq:on rew:off ctrs:Wt, Wff  use:renaming_up.
-  - move => n Γ a A B i hP ihP ha iha m Δ ξ hΔ hξ.
+  - move => Γ a A B i hP ihP ha iha Δ ξ hΔ hξ.
     apply : T_Abs; eauto.
     move : ihP(hΔ) (hξ); repeat move/[apply]. move/Bind_Inv.
     hauto lq:on rew:off ctrs:Wff,Wt use:renaming_up.
   - move => *. apply : T_App'; eauto. by asimpl.
-  - move => n Γ a A b B i hA ihA hB ihB hS ihS m Δ ξ hξ hΔ.
+  - move => Γ a A b B i hA ihA hB ihB hS ihS Δ ξ hξ hΔ.
     eapply T_Pair' with (U := ren_PTm ξ (subst_PTm (scons a VarPTm) B));eauto. by asimpl.
-  - move => n Γ a A B ha iha m Δ ξ hΔ hξ. apply : T_Proj2'; eauto. by asimpl.
-  - move => s Γ P a b c i hP ihP ha iha hb ihb hc ihc m Δ ξ hΔ hξ.
+  - move => Γ a A B ha iha Δ ξ hΔ hξ. apply : T_Proj2'; eauto. by asimpl.
+  - move => Γ P a b c i hP ihP ha iha hb ihb hc ihc Δ ξ hΔ hξ.
     move => [:hP].
     apply : T_Ind'; eauto. by asimpl.
     abstract :hP. apply ihP. by eauto using Wff_Cons', T_Nat'.
     hauto l:on use:renaming_up.
     set x := subst_PTm _ _. have -> : x = ren_PTm ξ (subst_PTm (scons PZero VarPTm) P) by subst x; asimpl.
     by subst x; eauto.
-    set Ξ := funcomp _ _.
+    set Ξ := cons _ _.
     have  hΞ : ⊢ Ξ. subst Ξ.
     apply : Wff_Cons'; eauto. apply hP.
-    move /(_ _ Ξ (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
-    set Ξ' := (funcomp _ _ : fin (S (S _)) -> _) .
+    move /(_ Ξ (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
+    set Ξ' := (cons _ _) .
     have : renaming_ok Ξ Ξ' (upRen_PTm_PTm (upRen_PTm_PTm ξ)).
     by do 2 apply renaming_up.
     move /[swap] /[apply].
@@ -225,31 +222,33 @@ Proof.
   - hauto lq:on ctrs:Wt,LEq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on rew:off use:E_Bind', Wff_Cons, renaming_up.
-  - move => n Γ a b A B i hPi ihPi ha iha m Δ ξ hΔ hξ.
+  - move => Γ a b A B i hPi ihPi ha iha Δ ξ hΔ hξ.
     move : ihPi (hΔ) (hξ). repeat move/[apply].
     move => /Bind_Inv [j][h0][h1]h2.
     have ? : Δ ⊢ PBind PPi (ren_PTm ξ A) (ren_PTm (upRen_PTm_PTm ξ) B) ∈ PUniv j by qauto l:on ctrs:Wt.
     move {hPi}.
     apply : E_Abs; eauto. qauto l:on ctrs:Wff use:renaming_up.
   - move => *. apply : E_App'; eauto. by asimpl.
-  - move => n Γ a0 a1 b0 b1 A B i hA ihA ha iha hb ihb m Δ ξ hΔ hξ.
+  - move => Γ a0 a1 b0 b1 A B i hA ihA ha iha hb ihb Δ ξ hΔ hξ.
     apply : E_Pair; eauto.
     move : ihb hΔ hξ. repeat move/[apply].
     by asimpl.
   - move => *. apply : E_Proj2'; eauto. by asimpl.
-  - move => n Γ P0 P1 a0 a1 b0 b1 c0 c1 i hP0 ihP0 hP ihP ha iha hb ihb hc ihc.
-    move => m Δ ξ hΔ hξ [:hP'].
-    apply E_IndCong' with (i := i).
+  - move => Γ P0 P1 a0 a1 b0 b1 c0 c1 i hP0 ihP0 hP ihP ha iha hb ihb hc ihc.
+    move => Δ ξ hΔ hξ [:hP' hren].
+    apply E_IndCong' with (i := i); eauto with wt.
     by asimpl.
+    apply ihP0.
     abstract : hP'.
     qauto l:on use:renaming_up, Wff_Cons', T_Nat'.
+    abstract : hren.
     qauto l:on use:renaming_up, Wff_Cons', T_Nat'.
-    by eauto with wt.
+    apply ihP; eauto with wt.
     move : ihb (hΔ) (hξ); do! move/[apply]. by asimpl.
-    set Ξ := funcomp _ _.
+    set Ξ := cons _ _.
     have hΞ : ⊢ Ξ.
-    subst Ξ. apply :Wff_Cons'; eauto. apply hP'.
-    move /(_ _ Ξ  (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
+    subst Ξ. apply :Wff_Cons'; eauto.
+    move /(_ Ξ  (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
     move /(_ ltac:(qauto l:on use:renaming_up)).
     suff : ren_PTm (upRen_PTm_PTm (upRen_PTm_PTm ξ))
       (ren_PTm shift
@@ -259,31 +258,33 @@ Proof.
          (ren_PTm (upRen_PTm_PTm ξ) P0)) by scongruence.
     by asimpl.
   - qauto l:on ctrs:Eq, LEq.
-  - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ξ hΔ hξ.
+  - move => Γ a b A B i hP ihP hb ihb ha iha Δ ξ hΔ hξ.
     move : ihP (hξ) (hΔ). repeat move/[apply].
     move /Bind_Inv.
     move => [j][h0][h1]h2.
     have ? : Δ ⊢ PBind PPi (ren_PTm ξ A) (ren_PTm (upRen_PTm_PTm ξ) B) ∈ PUniv j by qauto l:on ctrs:Wt.
     apply : E_AppAbs'; eauto. by asimpl. by asimpl.
     hauto lq:on ctrs:Wff use:renaming_up.
-  - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ξ hΔ hξ.
+  - move => Γ a b A B i hP ihP ha iha hb ihb Δ ξ hΔ hξ.
     move : {hP} ihP (hξ) (hΔ). repeat move/[apply].
     move /Bind_Inv => [i0][h0][h1]h2.
     have ? : Δ ⊢ PBind PSig (ren_PTm ξ A) (ren_PTm (upRen_PTm_PTm ξ) B) ∈ PUniv i0 by qauto l:on ctrs:Wt.
     apply : E_ProjPair1; eauto.
     move : ihb hξ hΔ. repeat move/[apply]. by asimpl.
-  - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ξ hΔ hξ.
+  - move => Γ a b A B i hP ihP ha iha hb ihb Δ ξ hΔ hξ.
     apply : E_ProjPair2'; eauto. by asimpl.
     move : ihb hξ hΔ; repeat move/[apply]. by asimpl.
-  - move => n Γ P i b c hP ihP hb ihb hc ihc m Δ ξ hΔ hξ.
+  - move => Γ P i b c hP ihP hb ihb hc ihc Δ ξ hΔ hξ.
     apply E_IndZero' with (i := i); eauto. by asimpl.
-    qauto l:on use:Wff_Cons', T_Nat', renaming_up.
-    move  /(_ m Δ ξ hΔ) : ihb.
+    sauto lq:on use:Wff_Cons', T_Nat', renaming_up.
+    move  /(_ Δ ξ hΔ) : ihb.
     asimpl. by apply.
-    set Ξ := funcomp _ _.
-    have hΞ : ⊢ Ξ by  qauto l:on use:Wff_Cons', T_Nat', renaming_up.
-    move /(_ _ Ξ  (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
-    move /(_ ltac:(qauto l:on use:renaming_up)).
+    set Ξ := cons _ _.
+    have hΞ : ⊢ Ξ by  sauto lq:on use:Wff_Cons', T_Nat', renaming_up.
+    move /(_ Ξ  (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
+    set p := renaming_ok _ _ _.
+    have : p. by do 2 apply renaming_up.
+    move => /[swap]/[apply].
     suff : ren_PTm (upRen_PTm_PTm (upRen_PTm_PTm ξ))
       (ren_PTm shift
          (subst_PTm
@@ -291,15 +292,15 @@ Proof.
       (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift))
          (ren_PTm (upRen_PTm_PTm ξ) P)) by scongruence.
     by asimpl.
-  - move => n Γ P a b c i hP ihP ha iha hb ihb hc ihc m Δ ξ hΔ hξ.
+  - move => Γ P a b c i hP ihP ha iha hb ihb hc ihc Δ ξ hΔ hξ.
     apply E_IndSuc' with (i := i). by asimpl. by asimpl.
-    qauto l:on use:Wff_Cons', T_Nat', renaming_up.
+    sauto lq:on use:Wff_Cons', T_Nat', renaming_up.
     by eauto with wt.
-    move  /(_ m Δ ξ hΔ) : ihb. asimpl. by apply.
-    set Ξ := funcomp _ _.
-    have hΞ : ⊢ Ξ by  qauto l:on use:Wff_Cons', T_Nat', renaming_up.
-    move /(_ _ Ξ  (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
-    move /(_ ltac:(qauto l:on use:renaming_up)).
+    move  /(_ Δ ξ hΔ) : ihb. asimpl. by apply.
+    set Ξ := cons _ _.
+    have hΞ : ⊢ Ξ by  sauto lq:on use:Wff_Cons', T_Nat', renaming_up.
+    move /(_ Ξ  (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
+    move /(_ ltac:(by eauto using renaming_up)).
     by asimpl.
   - move => *.
     apply : E_AppEta'; eauto. by asimpl.
@@ -315,94 +316,94 @@ Proof.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
 Qed.
 
-Definition morphing_ok {n m} Δ Γ (ρ : fin n -> PTm m) :=
-  forall i, Δ ⊢ ρ i ∈ subst_PTm ρ (Γ i).
+Definition morphing_ok Δ Γ (ρ : nat -> PTm) :=
+  forall i A, lookup i Γ A -> Δ ⊢ ρ i ∈ subst_PTm ρ A.
 
-Lemma morphing_ren n m p Ξ Δ Γ
-  (ρ : fin n -> PTm m) (ξ : fin m -> fin p) :
+Lemma morphing_ren Ξ Δ Γ
+  (ρ : nat -> PTm) (ξ : nat -> nat) :
   ⊢ Ξ ->
   renaming_ok Ξ Δ ξ -> morphing_ok Δ Γ ρ ->
   morphing_ok Ξ Γ (funcomp (ren_PTm ξ) ρ).
 Proof.
-  move => hΞ hξ hρ.
-  move => i.
+  move => hΞ hξ hρ i A.
   rewrite {1}/funcomp.
   have -> :
-    subst_PTm (funcomp (ren_PTm ξ) ρ) (Γ i) =
-    ren_PTm ξ (subst_PTm ρ (Γ i)) by asimpl.
-  eapply renaming; eauto.
+    subst_PTm (funcomp (ren_PTm ξ) ρ) A =
+    ren_PTm ξ (subst_PTm ρ A) by asimpl.
+  move => ?.  eapply renaming; eauto.
 Qed.
 
-Lemma morphing_ext n m Δ Γ (ρ : fin n -> PTm m) (a : PTm m) (A : PTm n)  :
+Lemma morphing_ext Δ Γ (ρ : nat -> PTm) (a : PTm) (A : PTm)  :
   morphing_ok Δ Γ ρ ->
   Δ ⊢ a ∈ subst_PTm ρ A ->
   morphing_ok
-  Δ (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ).
+    Δ (cons A Γ) (scons a ρ).
 Proof.
-  move => h ha i. destruct i as [i|]; by asimpl.
+  move => h ha i B.
+  elim /lookup_inv => //=_.
+  - move => A0 Γ0 ? [*]. subst.
+    by asimpl.
+  - move => i0 Γ0 A0 B0 h0 ? [*]. subst.
+    asimpl. by apply h.
 Qed.
 
-Lemma T_Var' n Γ (i : fin n) U :
-  U = Γ i ->
-  ⊢ Γ ->
-  Γ ⊢ VarPTm i ∈ U.
-Proof. move => ->. apply T_Var. Qed.
-
-Lemma renaming_wt : forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> forall m Δ (ξ : fin n -> fin m), ⊢ Δ -> renaming_ok Δ Γ ξ -> Δ ⊢ ren_PTm ξ a ∈ ren_PTm ξ A.
+Lemma renaming_wt : forall Γ (a A : PTm), Γ ⊢ a ∈ A -> forall Δ (ξ : nat -> nat), ⊢ Δ -> renaming_ok Δ Γ ξ -> Δ ⊢ ren_PTm ξ a ∈ ren_PTm ξ A.
 Proof. sfirstorder use:renaming. Qed.
 
-Lemma renaming_wt' : forall n m Δ Γ a A (ξ : fin n -> fin m) u U,
+Lemma renaming_wt' : forall Δ Γ a A (ξ : nat -> nat) u U,
     u = ren_PTm ξ a -> U = ren_PTm ξ A ->
     Γ ⊢ a ∈ A -> ⊢ Δ ->
     renaming_ok Δ Γ ξ -> Δ ⊢ u ∈ U.
 Proof. hauto use:renaming_wt. Qed.
 
-Lemma morphing_up n m Γ Δ (ρ : fin n -> PTm m) (A : PTm n) k :
+Lemma morphing_up Γ Δ (ρ : nat -> PTm) (A : PTm) k :
   morphing_ok Γ Δ ρ ->
   Γ ⊢ subst_PTm ρ A ∈ PUniv k ->
-  morphing_ok (funcomp (ren_PTm shift) (scons (subst_PTm ρ A) Γ)) (funcomp (ren_PTm shift) (scons A Δ)) (up_PTm_PTm ρ).
+  morphing_ok (cons (subst_PTm ρ A) Γ) (cons A Δ) (up_PTm_PTm ρ).
 Proof.
   move => h h1 [:hp]. apply morphing_ext.
   rewrite /morphing_ok.
-  move => i.
-  rewrite {2}/funcomp.
-  apply : renaming_wt'; eauto. by asimpl.
+  move => i A0 hA0.
+  apply : renaming_wt'; last by apply renaming_shift.
+  rewrite /funcomp. reflexivity.
+  2 : { eauto. }
+  by asimpl.
   abstract : hp. qauto l:on ctrs:Wff use:wff_mutual.
-  eauto using renaming_shift.
-  apply : T_Var';eauto. rewrite /funcomp. by asimpl.
+  apply : T_Var;eauto. apply here'. by asimpl.
 Qed.
 
-Lemma weakening_wt : forall n Γ (a A B : PTm n) i,
+Lemma weakening_wt : forall Γ (a A B : PTm) i,
     Γ ⊢ B ∈ PUniv i ->
     Γ ⊢ a ∈ A ->
-    funcomp (ren_PTm shift) (scons B Γ) ⊢ ren_PTm shift a ∈ ren_PTm shift A.
+    cons B Γ ⊢ ren_PTm shift a ∈ ren_PTm shift A.
 Proof.
-  move => n Γ a A B i hB ha.
+  move => Γ a A B i hB ha.
   apply : renaming_wt'; eauto.
   apply : Wff_Cons'; eauto.
   apply : renaming_shift; eauto.
 Qed.
 
-Lemma weakening_wt' : forall n Γ (a A B : PTm n) i U u,
+Lemma weakening_wt' : forall Γ (a A B : PTm) i U u,
     u = ren_PTm shift a ->
     U = ren_PTm shift A ->
     Γ ⊢ B ∈ PUniv i ->
     Γ ⊢ a ∈ A ->
-    funcomp (ren_PTm shift) (scons B Γ) ⊢ u ∈ U.
+    cons B Γ ⊢ u ∈ U.
 Proof. move => > -> ->. apply weakening_wt. Qed.
 
 Lemma morphing :
-  (forall n (Γ : fin n -> PTm n), ⊢ Γ -> True) /\
-  (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> forall m Δ (ρ : fin n -> PTm m), ⊢ Δ -> morphing_ok Δ Γ ρ ->
+  (forall Γ, ⊢ Γ -> True) /\
+    (forall Γ a A, Γ ⊢ a ∈ A -> forall Δ ρ, ⊢ Δ -> morphing_ok Δ Γ ρ ->
      Δ ⊢ subst_PTm ρ a ∈ subst_PTm ρ A) /\
-  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> forall m Δ (ρ : fin n -> PTm m), ⊢ Δ -> morphing_ok Δ Γ ρ ->
+    (forall Γ a b A, Γ ⊢ a ≡ b ∈ A -> forall Δ ρ, ⊢ Δ -> morphing_ok Δ Γ ρ ->
      Δ ⊢ subst_PTm ρ a ≡ subst_PTm ρ b ∈ subst_PTm ρ A) /\
-  (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> forall  m Δ (ρ : fin n -> PTm m), ⊢ Δ -> morphing_ok Δ Γ ρ ->
+    (forall Γ A B, Γ ⊢ A ≲ B -> forall Δ ρ, ⊢ Δ -> morphing_ok Δ Γ ρ ->
     Δ ⊢ subst_PTm ρ A ≲ subst_PTm ρ B).
 Proof.
   apply wt_mutual => //=.
+  - hauto l:on unfold:morphing_ok.
   - hauto lq:on use:morphing_up, Wff_Cons', T_Bind.
-  - move => n Γ a A B i hP ihP ha iha m Δ ρ hΔ hρ.
+  - move => Γ a A B i hP ihP ha iha Δ ρ hΔ hρ.
     move : ihP (hΔ) (hρ); repeat move/[apply].
     move /Bind_Inv => [j][h0][h1]h2. move {hP}.
     have ? : Δ ⊢ PBind PPi (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) ∈ PUniv i by hauto lq:on ctrs:Wt.
@@ -411,7 +412,7 @@ Proof.
     hauto lq:on use:Wff_Cons', Bind_Inv.
     apply : morphing_up; eauto.
   - move => *; apply : T_App'; eauto; by asimpl.
-  - move => n Γ a A b B i hA ihA hB ihB hS ihS m Δ ρ hρ hΔ.
+  - move => Γ a A b B i hA ihA hB ihB hS ihS Δ ρ hρ hΔ.
     eapply T_Pair' with (U := subst_PTm ρ (subst_PTm (scons a VarPTm) B));eauto. by asimpl.
   - hauto lq:on use:T_Proj1.
   - move => *. apply : T_Proj2'; eauto. by asimpl.
@@ -419,9 +420,8 @@ Proof.
   - qauto l:on ctrs:Wt.
   - qauto l:on ctrs:Wt.
   - qauto l:on ctrs:Wt.
-  - move => s Γ P a b c i hP ihP ha iha hb ihb hc ihc m Δ ξ hΔ hξ.
-    have  hξ' : morphing_ok (funcomp (ren_PTm shift) (scons PNat Δ))
-    (funcomp (ren_PTm shift) (scons PNat Γ)) (up_PTm_PTm ξ).
+  - move => Γ P a b c i hP ihP ha iha hb ihb hc ihc Δ ξ hΔ hξ.
+    have  hξ' : morphing_ok (cons PNat Δ) (cons PNat Γ) (up_PTm_PTm ξ).
     move /morphing_up : hξ.  move /(_ PNat 0).
     apply. by apply T_Nat.
     move => [:hP].
@@ -430,11 +430,11 @@ Proof.
     move /morphing_up : hξ.  move /(_ PNat 0).
     apply. by apply T_Nat.
     move : ihb hξ hΔ; do!move/[apply]. by asimpl.
-    set Ξ := funcomp _ _.
+    set Ξ := cons _ _.
     have  hΞ : ⊢ Ξ. subst Ξ.
     apply : Wff_Cons'; eauto. apply hP.
-    move /(_ _ Ξ (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
-    set Ξ' := (funcomp _ _ : fin (S (S _)) -> _) .
+    move /(_ Ξ (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    set Ξ' := (cons _ _) .
     have : morphing_ok Ξ Ξ' (up_PTm_PTm (up_PTm_PTm ξ)).
     eapply morphing_up; eauto. apply hP.
     move /[swap]/[apply].
@@ -447,23 +447,23 @@ Proof.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on rew:off use:E_Bind', Wff_Cons, morphing_up.
-  - move => n Γ a b A B i hPi ihPi ha iha m Δ ρ hΔ hρ.
+  - move => Γ a b A B i hPi ihPi ha iha Δ ρ hΔ hρ.
     move : ihPi (hΔ) (hρ). repeat move/[apply].
     move => /Bind_Inv [j][h0][h1]h2.
     have ? : Δ ⊢ PBind PPi (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) ∈ PUniv j by qauto l:on ctrs:Wt.
     move {hPi}.
     apply : E_Abs; eauto. qauto l:on ctrs:Wff use:morphing_up.
   - move => *. apply : E_App'; eauto. by asimpl.
-  - move => n Γ a0 a1 b0 b1 A B i hA ihA ha iha hb ihb m Δ ρ hΔ hρ.
+  - move => Γ a0 a1 b0 b1 A B i hA ihA ha iha hb ihb Δ ρ hΔ hρ.
     apply : E_Pair; eauto.
     move : ihb hΔ hρ. repeat move/[apply].
     by asimpl.
   - hauto q:on ctrs:Eq.
   - move => *. apply : E_Proj2'; eauto. by asimpl.
-  - move => n Γ P0 P1 a0 a1 b0 b1 c0 c1 i hP0 ihP0 hP ihP ha iha hb ihb hc ihc.
-    move => m Δ ξ hΔ hξ.
-    have  hξ' : morphing_ok (funcomp (ren_PTm shift) (scons PNat Δ))
-                  (funcomp (ren_PTm shift) (scons PNat Γ)) (up_PTm_PTm ξ).
+  - move => Γ P0 P1 a0 a1 b0 b1 c0 c1 i hP0 ihP0 hP ihP ha iha hb ihb hc ihc.
+    move => Δ ξ hΔ hξ.
+    have  hξ' : morphing_ok (cons PNat Δ)
+                   (cons PNat Γ) (up_PTm_PTm ξ).
     move /morphing_up : hξ.  move /(_ PNat 0).
     apply. by apply T_Nat.
     move => [:hP'].
@@ -474,56 +474,54 @@ Proof.
     qauto l:on use:renaming_up, Wff_Cons', T_Nat'.
     by eauto with wt.
     move : ihb (hΔ) (hξ); do! move/[apply]. by asimpl.
-    set Ξ := funcomp _ _.
+    set Ξ := cons _ _.
     have hΞ : ⊢ Ξ.
     subst Ξ. apply :Wff_Cons'; eauto. apply hP'.
-    move /(_ _ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    move /(_ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
     move /(_ ltac:(qauto l:on use:morphing_up)).
     asimpl. substify. by asimpl.
   - eauto with wt.
   - qauto l:on ctrs:Eq, LEq.
-  - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ρ hΔ hρ.
+  - move => Γ a b A B i hP ihP hb ihb ha iha Δ ρ hΔ hρ.
     move : ihP (hρ) (hΔ). repeat move/[apply].
     move /Bind_Inv.
     move => [j][h0][h1]h2.
     have ? : Δ ⊢ PBind PPi (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) ∈ PUniv j by qauto l:on ctrs:Wt.
     apply : E_AppAbs'; eauto. by asimpl. by asimpl.
     hauto lq:on ctrs:Wff use:morphing_up.
-  - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ρ hΔ hρ.
+  - move => Γ a b A B i hP ihP ha iha hb ihb Δ ρ hΔ hρ.
     move : {hP} ihP (hρ) (hΔ). repeat move/[apply].
     move /Bind_Inv => [i0][h0][h1]h2.
     have ? : Δ ⊢ PBind PSig (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) ∈ PUniv i0 by qauto l:on ctrs:Wt.
     apply : E_ProjPair1; eauto.
     move : ihb hρ hΔ. repeat move/[apply]. by asimpl.
-  - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ρ hΔ hρ.
+  - move => Γ a b A B i hP ihP ha iha hb ihb Δ ρ hΔ hρ.
     apply : E_ProjPair2'; eauto. by asimpl.
     move : ihb hρ hΔ; repeat move/[apply]. by asimpl.
-  - move => n Γ P i b c hP ihP hb ihb hc ihc m Δ ξ hΔ hξ.
-    have  hξ' : morphing_ok (funcomp (ren_PTm shift) (scons PNat Δ))
-                  (funcomp (ren_PTm shift) (scons PNat Γ)) (up_PTm_PTm ξ).
+  - move => Γ P i b c hP ihP hb ihb hc ihc Δ ξ hΔ hξ.
+    have  hξ' : morphing_ok (cons PNat Δ)(cons PNat Γ) (up_PTm_PTm ξ).
     move /morphing_up : hξ.  move /(_ PNat 0).
     apply. by apply T_Nat.
     apply E_IndZero' with (i := i); eauto. by asimpl.
     qauto l:on use:Wff_Cons', T_Nat', renaming_up.
-    move  /(_ m Δ ξ hΔ) : ihb.
+    move  /(_ Δ ξ hΔ) : ihb.
     asimpl. by apply.
-    set Ξ := funcomp _ _.
+    set Ξ := cons _ _.
     have hΞ : ⊢ Ξ by  qauto l:on use:Wff_Cons', T_Nat', renaming_up.
-    move /(_ _ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    move /(_ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
     move /(_ ltac:(sauto lq:on use:morphing_up)).
     asimpl. substify. by asimpl.
-  - move => n Γ P a b c i hP ihP ha iha hb ihb hc ihc m Δ ξ hΔ hξ.
-    have  hξ' : morphing_ok (funcomp (ren_PTm shift) (scons PNat Δ))
-                  (funcomp (ren_PTm shift) (scons PNat Γ)) (up_PTm_PTm ξ).
+  - move => Γ P a b c i hP ihP ha iha hb ihb hc ihc Δ ξ hΔ hξ.
+    have  hξ' : morphing_ok (cons PNat Δ) (cons PNat Γ) (up_PTm_PTm ξ).
     move /morphing_up : hξ.  move /(_ PNat 0).
     apply. by apply T_Nat'.
     apply E_IndSuc' with (i := i). by asimpl. by asimpl.
     qauto l:on use:Wff_Cons', T_Nat', renaming_up.
     by eauto with wt.
-    move  /(_ m Δ ξ hΔ) : ihb. asimpl. by apply.
-    set Ξ := funcomp _ _.
+    move  /(_ Δ ξ hΔ) : ihb. asimpl. by apply.
+    set Ξ := cons _ _.
     have hΞ : ⊢ Ξ by  qauto l:on use:Wff_Cons', T_Nat', renaming_up.
-    move /(_ _ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    move /(_ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
     move /(_ ltac:(sauto l:on use:morphing_up)).
     asimpl. substify. by asimpl.
   - move => *.
@@ -540,40 +538,39 @@ Proof.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
 Qed.
 
-Lemma morphing_wt : forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> forall m Δ (ρ : fin n -> PTm m), ⊢ Δ -> morphing_ok Δ Γ ρ -> Δ ⊢ subst_PTm ρ a ∈ subst_PTm ρ A.
+Lemma morphing_wt : forall Γ (a A : PTm ), Γ ⊢ a ∈ A -> forall Δ (ρ : nat -> PTm), ⊢ Δ -> morphing_ok Δ Γ ρ -> Δ ⊢ subst_PTm ρ a ∈ subst_PTm ρ A.
 Proof. sfirstorder use:morphing. Qed.
 
-Lemma morphing_wt' : forall n m Δ Γ a A (ρ : fin n -> PTm m) u U,
+Lemma morphing_wt' : forall Δ Γ a A (ρ : nat -> PTm) u U,
     u = subst_PTm ρ a -> U = subst_PTm ρ A ->
     Γ ⊢ a ∈ A -> ⊢ Δ ->
     morphing_ok Δ Γ ρ -> Δ ⊢ u ∈ U.
 Proof. hauto use:morphing_wt. Qed.
 
-Lemma morphing_id : forall n (Γ : fin n -> PTm n), ⊢ Γ -> morphing_ok Γ Γ VarPTm.
+Lemma morphing_id : forall Γ, ⊢ Γ -> morphing_ok Γ Γ VarPTm.
 Proof.
-  move => n Γ hΓ.
-  rewrite /morphing_ok.
-  move => i. asimpl. by apply T_Var.
+  rewrite /morphing_ok => Γ hΓ i. asimpl.
+  eauto using T_Var.
 Qed.
 
-Lemma substing_wt : forall n Γ (a : PTm (S n)) (b A : PTm n) B,
-    funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
+Lemma substing_wt : forall Γ (a : PTm) (b A : PTm) B,
+    (cons A Γ) ⊢ a ∈ B ->
     Γ ⊢ b ∈ A ->
     Γ ⊢ subst_PTm (scons b VarPTm) a ∈ subst_PTm (scons b VarPTm) B.
 Proof.
-  move => n Γ a b A B ha hb [:hΓ]. apply : morphing_wt; eauto.
+  move => Γ a b A B ha hb [:hΓ]. apply : morphing_wt; eauto.
   abstract : hΓ. sfirstorder use:wff_mutual.
   apply morphing_ext; last by asimpl.
   by apply morphing_id.
 Qed.
 
 (* Could generalize to all equal contexts *)
-Lemma ctx_eq_subst_one n (A0 A1 : PTm n) i j Γ a A :
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ a ∈ A ->
+Lemma ctx_eq_subst_one (A0 A1 : PTm) i j Γ a A :
+  (cons A0 Γ) ⊢ a ∈ A ->
   Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A1 ∈ PUniv j ->
   Γ ⊢ A1 ≲ A0 ->
-  funcomp (ren_PTm shift) (scons A1 Γ) ⊢ a ∈ A.
+  (cons A1 Γ) ⊢ a ∈ A.
 Proof.
   move => h0 h1 h2 h3.
   replace a with (subst_PTm VarPTm a); last by asimpl.
@@ -581,22 +578,23 @@ Proof.
   have ? : ⊢ Γ by sfirstorder use:wff_mutual.
   apply : morphing_wt; eauto.
   apply : Wff_Cons'; eauto.
-  move => k. destruct k as [k|].
-  - asimpl.
-    eapply weakening_wt' with (a := VarPTm k);eauto using T_Var.
+  move => k A2. elim/lookup_inv => //=_; cycle 1.
+  - move => i0 Γ0 A3 B hi ? [*]. subst.
+    asimpl.
+    eapply weakening_wt' with (a := VarPTm i0);eauto using T_Var.
     by substify.
-  - move => [:hΓ'].
+  - move => A3 Γ0 ? [*]. subst.
+    move => [:hΓ'].
     apply : T_Conv.
     apply T_Var.
     abstract : hΓ'.
-    eauto using Wff_Cons'.
-    rewrite /funcomp. asimpl. substify. asimpl.
-    renamify.
+    eauto using Wff_Cons'. apply here.
+    asimpl. renamify.
     eapply renaming; eauto.
     apply : renaming_shift; eauto.
 Qed.
 
-Lemma bind_inst n Γ p (A : PTm n) B i a0 a1 :
+Lemma bind_inst Γ p (A : PTm) B i a0 a1 :
   Γ ⊢ PBind p A B ∈ PUniv i ->
   Γ ⊢ a0 ≡ a1 ∈ A ->
   Γ ⊢ subst_PTm (scons a0 VarPTm) B ≲ subst_PTm (scons a1 VarPTm) B.
@@ -606,7 +604,7 @@ Proof.
   case : p h => //=; hauto l:on use:Su_Pi_Proj2, Su_Sig_Proj2.
 Qed.
 
-Lemma Cumulativity n Γ (a : PTm n) i j :
+Lemma Cumulativity Γ (a : PTm ) i j :
   i <= j ->
   Γ ⊢ a ∈ PUniv i ->
   Γ ⊢ a ∈ PUniv j.
@@ -616,9 +614,9 @@ Proof.
   sfirstorder use:wff_mutual.
 Qed.
 
-Lemma T_Bind' n Γ i j p (A : PTm n) (B : PTm (S n)) :
+Lemma T_Bind' Γ i j p (A : PTm ) (B : PTm) :
   Γ ⊢ A ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv j ->
+  (cons A Γ) ⊢ B ∈ PUniv j ->
   Γ ⊢ PBind p A B ∈ PUniv (max i j).
 Proof.
   move => h0 h1.
@@ -629,47 +627,48 @@ Qed.
 Hint Resolve T_Bind' : wt.
 
 Lemma regularity :
-  (forall n (Γ : fin n -> PTm n), ⊢ Γ -> forall i, exists j, Γ ⊢ Γ i ∈ PUniv j) /\
-  (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> exists i, Γ ⊢ A ∈ PUniv i)  /\
-  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> Γ ⊢ a ∈ A /\ Γ ⊢ b ∈ A /\ exists i, Γ ⊢ A ∈ PUniv i) /\
-  (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> exists i, Γ ⊢ A ∈ PUniv i /\ Γ ⊢ B ∈ PUniv i).
+  (forall Γ, ⊢ Γ -> forall i A, lookup i Γ A -> exists j, Γ ⊢ A ∈ PUniv j) /\
+    (forall Γ a A, Γ ⊢ a ∈ A -> exists i, Γ ⊢ A ∈ PUniv i)  /\
+    (forall Γ a b A, Γ ⊢ a ≡ b ∈ A -> Γ ⊢ a ∈ A /\ Γ ⊢ b ∈ A /\ exists i, Γ ⊢ A ∈ PUniv i) /\
+    (forall Γ A B, Γ ⊢ A ≲ B -> exists i, Γ ⊢ A ∈ PUniv i /\ Γ ⊢ B ∈ PUniv i).
 Proof.
   apply wt_mutual => //=; eauto with wt.
-  - move => n Γ A i hΓ ihΓ hA _ j.
-    destruct j as [j|].
-    have := ihΓ j.
-    move => [j0 hj].
-    exists j0. apply : renaming_wt' => //=; eauto using renaming_shift.
-    reflexivity. econstructor; eauto.
-    exists i. rewrite {2}/funcomp. simpl.
-    apply : renaming_wt'; eauto. reflexivity.
-    econstructor; eauto.
-    apply : renaming_shift; eauto.
-  - move => n Γ b a A B hb [i ihb] ha [j iha].
+  - move => i A. elim/lookup_inv => //=_.
+  - move => Γ A i hΓ ihΓ hA _ j A0.
+    elim /lookup_inv => //=_; cycle 1.
+    + move => i0 Γ0 A1 B + ? [*]. subst.
+      move : ihΓ => /[apply]. move => [j hA1].
+      exists j. apply : renaming_wt' => //=; eauto using renaming_shift. done.
+      apply : Wff_Cons'; eauto.
+    + move => A1 Γ0 ? [*]. subst.
+      exists i. apply : renaming_wt'; eauto. reflexivity.
+      econstructor; eauto.
+      apply : renaming_shift; eauto.
+  - move => Γ b a A B hb [i ihb] ha [j iha].
     move /Bind_Inv : ihb => [k][h0][h1]h2.
     move : substing_wt ha h1; repeat move/[apply].
     move => h. exists k.
     move : h. by asimpl.
   - hauto lq:on use:Bind_Inv.
-  - move => n Γ a A B ha [i /Bind_Inv[j][h0][h1]h2].
+  - move => Γ a A B ha [i /Bind_Inv[j][h0][h1]h2].
     exists j. have : Γ ⊢ PProj PL a ∈ A by qauto use:T_Proj1.
     move : substing_wt h1; repeat move/[apply].
     by asimpl.
-  - move => s Γ P a b c i + ? + *. clear. move => h ha. exists i.
+  - move => Γ P a b c i + ? + *. clear. move => h ha. exists i.
     hauto lq:on use:substing_wt.
   - sfirstorder.
   - sfirstorder.
   - sfirstorder.
-  - move => n Γ i p A0 A1 B0 B1 hΓ ihΓ hA0
+  - move => Γ i p A0 A1 B0 B1 hΓ ihΓ hA0
              [i0 ihA0] hA [ihA [ihA' [i1 ihA'']]].
-    move => hB [ihB0 [ihB1 [i2 ihB2]]].
     repeat split => //=.
     qauto use:T_Bind.
     apply T_Bind; eauto.
+    sfirstorder.
     apply : ctx_eq_subst_one; eauto using Su_Eq, E_Symmetric.
-    eauto using T_Univ.
+    hauto lq:on.
   - hauto lq:on ctrs:Wt,Eq.
-  - move => n Γ i b0 b1 a0 a1 A B hP _ hb [ihb0 [ihb1 [i0 ihb2]]]
+  - move => n i b0 b1 a0 a1 A B hP _ hb [ihb0 [ihb1 [i0 ihb2]]]
              ha [iha0 [iha1 [i1 iha2]]].
     repeat split.
     qauto use:T_App.
@@ -680,15 +679,16 @@ Proof.
     hauto lq:on ctrs:Wt,Eq,LEq lq:on use:Bind_Inv, substing_wt.
   - hauto lq:on use:bind_inst db:wt.
   - hauto lq:on use:Bind_Inv db:wt.
-  - move => n Γ i a b A B hS _ hab [iha][ihb][j]ihs.
+  - move => Γ i a b A B hS _ hab [iha][ihb][j]ihs.
     repeat split => //=; eauto with wt.
     apply : T_Conv; eauto with wt.
     move /E_Symmetric /E_Proj1 in hab.
     eauto using bind_inst.
     move /T_Proj1 in iha.
     hauto lq:on ctrs:Wt,Eq,LEq use:Bind_Inv, substing_wt.
-  - move => n Γ P0 P1 a0 a1 b0 b1 c0 c1 i _ _ hPE [hP0 [hP1 _]] hae [ha0 [ha1 _]] _ [hb0 [hb1 hb2]] _ [hc0 [hc1 hc2]].
+  - move => Γ P0 P1 a0 a1 b0 b1 c0 c1 i _ _ hPE [hP0 [hP1 _]] hae [ha0 [ha1 _]] _ [hb0 [hb1 hb2]] _ [hc0 [hc1 [j hc2]]].
     have wfΓ : ⊢ Γ by hauto use:wff_mutual.
+    have wfΓ' : ⊢ (PNat :: Γ) by qauto use:Wff_Cons', T_Nat'.
     repeat split. by eauto using T_Ind.
     apply : T_Conv. apply : T_Ind; eauto.
     apply : T_Conv; eauto.
@@ -696,13 +696,12 @@ Proof.
     apply : T_Conv.  apply : ctx_eq_subst_one; eauto.
     by eauto using Su_Eq, E_Symmetric.
     eapply renaming; eauto.
-    eapply morphing; eauto 20 using Su_Eq, morphing_ext, morphing_id with wt.
-    apply morphing_ext; eauto.
-    replace (funcomp VarPTm shift) with (funcomp (@ren_PTm n _ shift) VarPTm); last by asimpl.
-    apply : morphing_ren; eauto using Wff_Cons' with wt.
-    apply : renaming_shift; eauto. by apply morphing_id.
-    apply T_Suc. apply T_Var'. by asimpl. apply : Wff_Cons'; eauto using T_Nat'.
-    apply : Wff_Cons'; eauto. apply : renaming_shift; eauto.
+    eapply morphing. apply : Su_Eq. apply hPE. apply wfΓ'.
+    apply morphing_ext.
+    replace (funcomp VarPTm shift) with (funcomp (@ren_PTm shift) VarPTm); last by asimpl.
+    apply : morphing_ren; eauto using morphing_id, renaming_shift.
+    apply T_Suc. apply T_Var=>//. apply here. apply : Wff_Cons'; eauto using T_Nat'.
+    apply renaming_shift.
     have /E_Symmetric /Su_Eq : Γ ⊢ PBind PSig PNat P0 ≡ PBind PSig PNat P1 ∈ PUniv i by eauto with wt.
     move /E_Symmetric in hae.
     by eauto using Su_Sig_Proj2.
@@ -711,62 +710,62 @@ Proof.
   - hauto lq:on ctrs:Wt.
   - hauto q:on use:substing_wt db:wt.
   - hauto l:on use:bind_inst db:wt.
-  - move => n Γ P i b c hP _ hb _ hc _.
+  - move => Γ P i b c hP _ hb _ hc _.
     have ? : ⊢ Γ by hauto use:wff_mutual.
     repeat split=>//.
     by eauto with wt.
     sauto lq:on use:substing_wt.
-  - move => s Γ P a b c i hP _ ha _ hb _ hc _.
+  - move => Γ P a b c i hP _ ha _ hb _ hc _.
     have ? : ⊢ Γ by hauto use:wff_mutual.
     repeat split=>//.
     apply : T_Ind; eauto with wt.
-    set Ξ : fin (S (S _)) -> _ := (X in X ⊢ _ ∈ _) in hc.
+    set Ξ := (X in X ⊢ _ ∈ _) in hc.
     have : morphing_ok Γ Ξ (scons (PInd P a b c) (scons a VarPTm)).
     apply morphing_ext. apply morphing_ext. by apply morphing_id.
     by eauto. by eauto with wt.
     subst Ξ.
     move : morphing_wt hc; repeat move/[apply]. asimpl. by apply.
     sauto lq:on use:substing_wt.
-  - move => n Γ b A B i hΓ ihΓ hP _ hb [i0 ihb].
+  - move => Γ b A B i hP _ hb [i0 ihb].
     repeat split => //=; eauto with wt.
-    have {}hb : funcomp (ren_PTm shift) (scons A Γ) ⊢ ren_PTm shift b ∈ ren_PTm shift (PBind PPi A B)
+    have {}hb : (cons A Γ) ⊢ ren_PTm shift b ∈ ren_PTm shift (PBind PPi A B)
                         by hauto lq:on use:weakening_wt, Bind_Inv.
     apply : T_Abs; eauto.
-    apply : T_App'; eauto; rewrite-/ren_PTm.
-    by asimpl.
+    apply : T_App'; eauto; rewrite-/ren_PTm. asimpl. by rewrite subst_scons_id.
     apply T_Var. sfirstorder use:wff_mutual.
+    apply here.
   - hauto lq:on ctrs:Wt.
-  - move => n Γ A B C hA [i [ihA0 ihA1]] hC [j [ihC0 ihC1]].
+  - move => Γ A B C hA [i [ihA0 ihA1]] hC [j [ihC0 ihC1]].
     have ? : ⊢ Γ by sfirstorder use:wff_mutual.
     exists (max i j).
     have [? ?] : i <= Nat.max i j /\ j <= Nat.max i j by lia.
     qauto l:on use:T_Conv, Su_Univ.
-  - move => n Γ i j hΓ *. exists (S (max i j)).
+  - move => Γ i j hΓ *. exists (S (max i j)).
     have [? ?] : S i <= S (Nat.max i j) /\ S j <= S (Nat.max i j) by lia.
     hauto lq:on ctrs:Wt,LEq.
-  - move => n Γ A0 A1 B0 B1 i hΓ ihΓ hA0 _ hA1 [i0][ih0]ih1 hB[j0][ihB0]ihB1.
+  - move => Γ A0 A1 B0 B1 i hA0 _ hA1 [i0][ih0]ih1 hB[j0][ihB0]ihB1.
     exists (max i0 j0).
     split; eauto with wt.
     apply T_Bind'; eauto.
     sfirstorder use:ctx_eq_subst_one.
-  - move => n Γ A0 A1 B0 B1 i hΓ ihΓ hA1 _ hA0 [i0][ihA0]ihA1 hB[i1][ihB0]ihB1.
+  - move => n A0 A1 B0 B1 i hA1 _ hA0 [i0][ihA0]ihA1 hB[i1][ihB0]ihB1.
     exists (max i0 i1). repeat split; eauto with wt.
     apply T_Bind'; eauto.
     sfirstorder use:ctx_eq_subst_one.
   - sfirstorder.
-  - move => n Γ A0 A1 B0 B1 _ [i][ih0 ih1].
+  - move => Γ A0 A1 B0 B1 _ [i][ih0 ih1].
     move /Bind_Inv : ih0 => [i0][h _].
     move /Bind_Inv : ih1 => [i1][h' _].
     exists (max i0 i1).
     have [? ?] : i0 <= Nat.max i0 i1 /\ i1 <= Nat.max i0 i1 by lia.
     eauto using Cumulativity.
-  - move => n Γ A0 A1 B0 B1 _ [i][ih0 ih1].
+  - move => Γ A0 A1 B0 B1 _ [i][ih0 ih1].
     move /Bind_Inv : ih0 => [i0][h _].
     move /Bind_Inv : ih1 => [i1][h' _].
     exists (max i0 i1).
     have [? ?] : i0 <= Nat.max i0 i1 /\ i1 <= Nat.max i0 i1 by lia.
     eauto using Cumulativity.
-  - move => n Γ a0 a1 A0 A1 B0 B1 /Su_Pi_Proj1 hA1.
+  - move => Γ a0 a1 A0 A1 B0 B1 /Su_Pi_Proj1 hA1.
     move => [i][ihP0]ihP1.
     move => ha [iha0][iha1][j]ihA1.
     move /Bind_Inv :ihP0 => [i0][ih0][ih0' _].
@@ -779,7 +778,7 @@ Proof.
       move : substing_wt ih0';repeat move/[apply]. by asimpl.
     + apply Cumulativity with (i := i1); eauto.
       move : substing_wt ih1' iha1;repeat move/[apply]. by asimpl.
-  - move => n Γ a0 a1 A0 A1 B0 B1 /Su_Sig_Proj1 hA1.
+  - move => Γ a0 a1 A0 A1 B0 B1 /Su_Sig_Proj1 hA1.
     move => [i][ihP0]ihP1.
     move => ha [iha0][iha1][j]ihA1.
     move /Bind_Inv :ihP0 => [i0][ih0][ih0' _].
@@ -795,23 +794,25 @@ Proof.
       Unshelve. all: exact 0.
 Qed.
 
-Lemma Var_Inv n Γ (i : fin n) A :
+Lemma Var_Inv Γ (i : nat) B A :
   Γ ⊢ VarPTm i ∈ A ->
-  ⊢ Γ /\ Γ ⊢ Γ i ≲ A.
+  lookup i Γ B ->
+  ⊢ Γ /\ Γ ⊢ B ≲ A.
 Proof.
   move E : (VarPTm i) => u hu.
   move : i E.
-  elim : n Γ u A / hu=>//=.
-  - move => n Γ i hΓ i0 [?]. subst.
+  elim : Γ u A / hu=>//=.
+  - move => i Γ A hΓ hi i0 [?]. subst.
     repeat split => //=.
-    have h : Γ ⊢ VarPTm i ∈ Γ i by eauto using T_Var.
+    have ? : A = B  by eauto using lookup_deter. subst.
+    have h : Γ ⊢ VarPTm i ∈ B by eauto using T_Var.
     eapply regularity in h.
     move : h => [i0]?.
     apply : Su_Eq. apply E_Refl; eassumption.
   - sfirstorder use:Su_Transitive.
 Qed.
 
-Lemma renaming_su' : forall n m Δ Γ (ξ : fin n -> fin m) (A B : PTm n) u U ,
+Lemma renaming_su' : forall Δ Γ ξ (A B : PTm) u U ,
     u = ren_PTm ξ A ->
     U = ren_PTm ξ B ->
     Γ ⊢ A ≲ B ->
@@ -819,23 +820,23 @@ Lemma renaming_su' : forall n m Δ Γ (ξ : fin n -> fin m) (A B : PTm n) u U ,
     Δ ⊢ u ≲ U.
 Proof. move => > -> ->. hauto l:on use:renaming. Qed.
 
-Lemma weakening_su : forall n Γ (A0 A1 B : PTm n) i,
+Lemma weakening_su : forall Γ (A0 A1 B : PTm) i,
     Γ ⊢ B ∈ PUniv i ->
     Γ ⊢ A0 ≲ A1 ->
-    funcomp (ren_PTm shift) (scons B Γ) ⊢ ren_PTm shift A0 ≲ ren_PTm shift A1.
+    (cons B Γ) ⊢ ren_PTm shift A0 ≲ ren_PTm shift A1.
 Proof.
-  move => n Γ A0 A1 B i hB hlt.
+  move => Γ A0 A1 B i hB hlt.
   apply : renaming_su'; eauto.
   apply : Wff_Cons'; eauto.
   apply : renaming_shift; eauto.
 Qed.
 
-Lemma regularity_sub0 : forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> exists i, Γ ⊢ A ∈ PUniv i.
+Lemma regularity_sub0 : forall Γ (A B : PTm), Γ ⊢ A ≲ B -> exists i, Γ ⊢ A ∈ PUniv i.
 Proof. hauto lq:on use:regularity. Qed.
 
-Lemma Su_Pi_Proj2_Var n Γ (A0 A1 : PTm n) B0 B1 :
+Lemma Su_Pi_Proj2_Var Γ (A0 A1 : PTm) B0 B1 :
   Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
-  funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ≲ B1.
+  cons A1 Γ ⊢ B0 ≲ B1.
 Proof.
   move => h.
   have /Su_Pi_Proj1 h1 := h.
@@ -843,15 +844,16 @@ Proof.
   move /weakening_su : (h) h2. move => /[apply].
   move => h2.
   apply : Su_Pi_Proj2'; try eassumption; rewrite -?/ren_PTm; cycle 2.
-  apply E_Refl. apply T_Var' with (i := var_zero); eauto.
+  apply E_Refl. apply T_Var with (i := var_zero); eauto.
   sfirstorder use:wff_mutual.
-  by asimpl.
-  by asimpl.
+  apply here.
+  by asimpl; rewrite subst_scons_id.
+  by asimpl; rewrite subst_scons_id.
 Qed.
 
-Lemma Su_Sig_Proj2_Var n Γ (A0 A1 : PTm n) B0 B1 :
+Lemma Su_Sig_Proj2_Var Γ (A0 A1 : PTm) B0 B1 :
   Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≲ B1.
+  (cons A0 Γ) ⊢ B0 ≲ B1.
 Proof.
   move => h.
   have /Su_Sig_Proj1 h1 := h.
@@ -859,8 +861,9 @@ Proof.
   move /weakening_su : (h) h2. move => /[apply].
   move => h2.
   apply : Su_Sig_Proj2'; try eassumption; rewrite -?/ren_PTm; cycle 2.
-  apply E_Refl. apply T_Var' with (i := var_zero); eauto.
+  apply E_Refl. apply T_Var with (i := var_zero); eauto.
   sfirstorder use:wff_mutual.
-  by asimpl.
-  by asimpl.
+  apply here.
+  by asimpl; rewrite subst_scons_id.
+  by asimpl; rewrite subst_scons_id.
 Qed.

From fe418c2a785b2779456529d5dca74ce156c5bf1a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 15:29:50 -0500
Subject: [PATCH 09/24] Fix preservation and broken cases in logrel

---
 theories/admissible.v   | 43 +++++++++++------------------------------
 theories/logrel.v       |  5 +----
 theories/preservation.v | 40 +++++++++++++++++++-------------------
 theories/soundness.v    |  2 +-
 4 files changed, 33 insertions(+), 57 deletions(-)

diff --git a/theories/admissible.v b/theories/admissible.v
index 392e3cf..830ceec 100644
--- a/theories/admissible.v
+++ b/theories/admissible.v
@@ -1,45 +1,24 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing structural.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common typing structural.
 From Hammer Require Import Tactics.
 Require Import ssreflect.
 Require Import Psatz.
 Require Import Coq.Logic.FunctionalExtensionality.
 
-Derive Dependent Inversion wff_inv with (forall n (Γ : fin n -> PTm n), ⊢ Γ) Sort Prop.
+Derive Inversion wff_inv with (forall Γ, ⊢ Γ) Sort Prop.
 
-Lemma Wff_Cons_Inv n Γ (A : PTm n) :
-  ⊢ funcomp (ren_PTm shift) (scons A Γ) ->
-  ⊢ Γ /\ exists i, Γ ⊢ A ∈ PUniv i.
-Proof.
-  elim /wff_inv => //= _.
-  move => n0 Γ0 A0 i0 hΓ0 hA0 [? _]. subst.
-  Equality.simplify_dep_elim.
-  have h : forall i, (funcomp (ren_PTm shift) (scons A0 Γ0)) i = (funcomp (ren_PTm shift) (scons A Γ)) i by scongruence.
-  move /(_ var_zero) : (h).
-  rewrite /funcomp. asimpl.
-  move /ren_injective. move /(_ ltac:(hauto lq:on rew:off inv:option)).
-  move => ?. subst.
-  have : Γ0 = Γ. extensionality i.
-  move /(_ (Some i)) : h.
-  rewrite /funcomp. asimpl.
-  move /ren_injective. move /(_ ltac:(hauto lq:on rew:off inv:option)).
-  done.
-  move => ?. subst. sfirstorder.
-Qed.
-
-Lemma T_Abs n Γ (a : PTm (S n)) A B :
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
+Lemma T_Abs Γ (a : PTm ) A B :
+  (cons A Γ) ⊢ a ∈ B ->
   Γ ⊢ PAbs a ∈ PBind PPi A B.
 Proof.
   move => ha.
-  have [i hB] : exists i, funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i by sfirstorder use:regularity.
-  have hΓ : ⊢ funcomp (ren_PTm shift) (scons A Γ) by sfirstorder use:wff_mutual.
-  move /Wff_Cons_Inv : hΓ => [hΓ][j]hA.
-  hauto lq:on rew:off use:T_Bind', typing.T_Abs.
+  have [i hB] : exists i, (cons A Γ) ⊢ B ∈ PUniv i by sfirstorder use:regularity.
+  have hΓ : ⊢  (cons A Γ) by sfirstorder use:wff_mutual.
+  hauto lq:on rew:off inv:Wff use:T_Bind', typing.T_Abs.
 Qed.
 
-Lemma E_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
+Lemma E_Bind Γ i p (A0 A1 : PTm) B0 B1 :
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
+  (cons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i.
 Proof.
   move => h0 h1.
@@ -49,7 +28,7 @@ Proof.
   firstorder.
 Qed.
 
-Lemma E_App n Γ (b0 b1 a0 a1 : PTm n) A B :
+Lemma E_App Γ (b0 b1 a0 a1 : PTm ) A B :
   Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B ->
   Γ ⊢ a0 ≡ a1 ∈ A ->
   Γ ⊢ PApp b0 a0 ≡ PApp b1 a1 ∈ subst_PTm (scons a0 VarPTm) B.
@@ -59,7 +38,7 @@ Proof.
   move : E_App h. by repeat move/[apply].
 Qed.
 
-Lemma E_Proj2 n Γ (a b : PTm n) A B :
+Lemma E_Proj2 Γ (a b : PTm) A B :
   Γ ⊢ a ≡ b ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B.
 Proof.
diff --git a/theories/logrel.v b/theories/logrel.v
index ae4169c..a479f81 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1583,7 +1583,7 @@ Lemma SSu_Pi Γ (A0 A1 : PTm ) B0 B1 :
   cons A0 Γ ⊨ B0 ≲ B1 ->
   Γ ⊨ PBind PPi A0 B0 ≲ PBind PPi A1 B1.
 Proof.
-  move => hΓ hA hB.
+  move =>  hA hB.
   have ? : SN A0 /\ SN A1 /\ SN B0 /\ SN B1
     by hauto l:on use:SemLEq_SN_Sub.
   apply SemLEq_SemWt in hA, hB.
@@ -1592,7 +1592,6 @@ Proof.
   apply : SemWt_SemLEq; last by hauto l:on use:Sub.PiCong.
   hauto l:on use:ST_Bind'.
   apply ST_Bind'; eauto.
-  have hΓ' : ⊨ (cons A1 Γ) by hauto l:on use:SemWff_cons.
   move => ρ hρ.
   suff : ρ_ok (cons A0 Γ) ρ by hauto l:on.
   move : hρ. suff : Γ_sub' (A0 :: Γ) (A1 :: Γ)
@@ -1614,8 +1613,6 @@ Proof.
   apply : SemWt_SemLEq; last by hauto l:on use:Sub.SigCong.
   2 : { hauto l:on use:ST_Bind'. }
   apply ST_Bind'; eauto.
-  have hΓ' : ⊨ cons A1 Γ by hauto l:on use:SemWff_cons.
-  have hΓ'' : ⊨ cons A0 Γ by hauto l:on use:SemWff_cons.
   move => ρ hρ.
   suff : ρ_ok (cons A1 Γ) ρ by hauto l:on.
   move : hρ. suff : Γ_sub' (A1 :: Γ) (A0 :: Γ) by hauto l:on.
diff --git a/theories/preservation.v b/theories/preservation.v
index fe0a920..c8a2106 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -30,12 +30,12 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma Proj1_Inv n Γ (a : PTm n) U :
+Lemma Proj1_Inv Γ (a : PTm ) U :
   Γ ⊢ PProj PL a ∈ U ->
   exists A B, Γ ⊢ a ∈ PBind PSig A B /\ Γ ⊢ A ≲ U.
 Proof.
   move E : (PProj PL a) => u hu.
-  move :a E. elim : n Γ u U / hu => n //=.
+  move :a E. elim : Γ u U / hu => //=.
   - move => Γ a A B ha _ a0 [*]. subst.
     exists A, B. split => //=.
     eapply regularity in ha.
@@ -45,12 +45,12 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma Proj2_Inv n Γ (a : PTm n) U :
+Lemma Proj2_Inv Γ (a : PTm) U :
   Γ ⊢ PProj PR a ∈ U ->
   exists A B, Γ ⊢ a ∈ PBind PSig A B /\ Γ ⊢ subst_PTm (scons (PProj PL a) VarPTm) B ≲ U.
 Proof.
   move E : (PProj PR a) => u hu.
-  move :a E. elim : n Γ u U / hu => n //=.
+  move :a E. elim : Γ u U / hu => //=.
   - move => Γ a A B ha _ a0 [*]. subst.
     exists A, B. split => //=.
     have ha' := ha.
@@ -62,30 +62,30 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma Pair_Inv n Γ (a b : PTm n) U :
+Lemma Pair_Inv Γ (a b : PTm ) U :
   Γ ⊢ PPair a b ∈ U ->
   exists A B, Γ ⊢ a ∈ A /\
          Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B /\
          Γ ⊢ PBind PSig A B ≲ U.
 Proof.
   move E : (PPair a b) => u hu.
-  move : a b E. elim : n Γ u U / hu => n //=.
+  move : a b E. elim : Γ u U / hu => //=.
   - move => Γ a b A B i hS _ ha _ hb _ a0 b0 [*]. subst.
     exists A,B. repeat split => //=.
     move /E_Refl /Su_Eq : hS. apply.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma Ind_Inv n Γ P (a : PTm n) b c U :
+Lemma Ind_Inv Γ P (a : PTm) b c U :
   Γ ⊢ PInd P a b c ∈ U ->
-  exists i, funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i /\
+  exists i, (cons PNat Γ) ⊢ P ∈ PUniv i /\
        Γ ⊢ a ∈ PNat /\
        Γ ⊢ b ∈ subst_PTm (scons PZero VarPTm) P /\
-      funcomp (ren_PTm shift)(scons P (funcomp (ren_PTm shift) (scons PNat Γ))) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) /\
+          (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) /\
        Γ ⊢ subst_PTm (scons a VarPTm) P ≲ U.
 Proof.
   move E : (PInd P a b c)=> u hu.
-  move : P a b c E. elim : n Γ u U / hu => n //=.
+  move : P a b c E. elim : Γ u U / hu => //=.
   - move => Γ P a b c i hP _ ha _ hb _ hc _ P0 a0 b0 c0 [*]. subst.
     exists i. repeat split => //=.
     have : Γ ⊢ subst_PTm (scons a VarPTm) P ∈ subst_PTm (scons a VarPTm) (PUniv i) by hauto l:on use:substing_wt.
@@ -93,10 +93,10 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma E_AppAbs : forall n (a : PTm (S n)) (b : PTm n) (Γ : fin n -> PTm n) (A : PTm n),
+Lemma E_AppAbs : forall (a : PTm) (b : PTm) Γ (A : PTm),
   Γ ⊢ PApp (PAbs a) b ∈ A -> Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ A.
 Proof.
-  move => n a b Γ A ha.
+  move => a b Γ A ha.
   move /App_Inv : ha.
   move => [A0][B0][ha][hb]hS.
   move /Abs_Inv : ha => [A1][B1][ha]hS0.
@@ -112,10 +112,10 @@ Proof.
   eauto using T_Conv.
 Qed.
 
-Lemma E_ProjPair1 : forall n (a b : PTm n) (Γ : fin n -> PTm n) (A : PTm n),
+Lemma E_ProjPair1 : forall (a b : PTm) Γ (A : PTm),
     Γ ⊢ PProj PL (PPair a b) ∈ A -> Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A.
 Proof.
-  move => n a b Γ A.
+  move => a b Γ A.
   move /Proj1_Inv. move => [A0][B0][hab]hA0.
   move /Pair_Inv : hab => [A1][B1][ha][hb]hS.
   have [i ?]  : exists i, Γ ⊢ PBind PSig A1 B1 ∈ PUniv i by sfirstorder use:regularity_sub0.
@@ -125,25 +125,25 @@ Proof.
   apply : E_ProjPair1; eauto.
 Qed.
 
-Lemma Suc_Inv n Γ (a : PTm n) A :
+Lemma Suc_Inv Γ (a : PTm) A :
   Γ ⊢ PSuc a ∈ A -> Γ ⊢ a ∈ PNat /\ Γ ⊢ PNat ≲ A.
 Proof.
   move E : (PSuc a) => u hu.
   move : a E.
-  elim : n Γ u A /hu => //=.
-  - move => n Γ a ha iha a0 [*]. subst.
+  elim : Γ u A /hu => //=.
+  - move => Γ a ha iha a0 [*]. subst.
     split => //=. eapply wff_mutual in ha.
     apply : Su_Eq.
     apply E_Refl. by apply T_Nat'.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma RRed_Eq n Γ (a b : PTm n) A :
+Lemma RRed_Eq Γ (a b : PTm) A :
   Γ ⊢ a ∈ A ->
   RRed.R a b ->
   Γ ⊢ a ≡ b ∈ A.
 Proof.
-  move => + h. move : Γ A. elim : n a b /h => n.
+  move => + h. move : Γ A. elim : a b /h.
   - apply E_AppAbs.
   - move => p a b Γ A.
     case : p => //=.
@@ -214,7 +214,7 @@ Proof.
   - hauto lq:on use:Suc_Inv, E_Conv, E_SucCong.
 Qed.
 
-Theorem subject_reduction n Γ (a b A : PTm n) :
+Theorem subject_reduction Γ (a b A : PTm) :
   Γ ⊢ a ∈ A ->
   RRed.R a b ->
   Γ ⊢ b ∈ A.
diff --git a/theories/soundness.v b/theories/soundness.v
index b11dded..8bfeedf 100644
--- a/theories/soundness.v
+++ b/theories/soundness.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.fintype Autosubst2.syntax.
+Require Import Autosubst2.unscoped Autosubst2.syntax.
 Require Import fp_red logrel typing.
 From Hammer Require Import Tactics.
 

From 3e8ad2c5bcb34072671da791b9abd15b7e83627f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 15:50:08 -0500
Subject: [PATCH 10/24] Work on the refactoring proof

---
 theories/algorithmic.v | 249 +++++++++++++++++++----------------------
 theories/soundness.v   |  10 +-
 theories/structural.v  |  13 +--
 3 files changed, 127 insertions(+), 145 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 096ec94..a52ac7e 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax
   common typing preservation admissible fp_red structural soundness.
 From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
@@ -8,7 +8,7 @@ Require Import Coq.Logic.FunctionalExtensionality.
 Require Import Cdcl.Itauto.
 
 Module HRed.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
+  Inductive R : PTm -> PTm ->  Prop :=
   (****************** Beta ***********************)
   | AppAbs a b :
     R (PApp (PAbs a) b)  (subst_PTm (scons b VarPTm) a)
@@ -33,46 +33,46 @@ Module HRed.
     R a0 a1 ->
     R (PInd P a0 b c) (PInd P a1 b c).
 
-  Lemma ToRRed n (a b : PTm n) : HRed.R a b -> RRed.R a b.
+  Lemma ToRRed (a b : PTm) : HRed.R a b -> RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
 
-  Lemma preservation n Γ (a b A : PTm n) : Γ ⊢ a ∈ A -> HRed.R a b -> Γ ⊢ b ∈ A.
+  Lemma preservation Γ (a b A : PTm) : Γ ⊢ a ∈ A -> HRed.R a b -> Γ ⊢ b ∈ A.
   Proof.
     sfirstorder use:subject_reduction, ToRRed.
   Qed.
 
-  Lemma ToEq n Γ (a b : PTm n) A : Γ ⊢ a ∈ A -> HRed.R a b -> Γ ⊢ a ≡ b ∈ A.
+  Lemma ToEq Γ (a b : PTm ) A : Γ ⊢ a ∈ A -> HRed.R a b -> Γ ⊢ a ≡ b ∈ A.
   Proof. sfirstorder use:ToRRed, RRed_Eq. Qed.
 
 End HRed.
 
 Module HReds.
-  Lemma preservation n Γ (a b A : PTm n) : Γ ⊢ a ∈ A -> rtc HRed.R a b -> Γ ⊢ b ∈ A.
+  Lemma preservation Γ (a b A : PTm) : Γ ⊢ a ∈ A -> rtc HRed.R a b -> Γ ⊢ b ∈ A.
   Proof. induction 2; sfirstorder use:HRed.preservation. Qed.
 
-  Lemma ToEq n Γ (a b : PTm n) A : Γ ⊢ a ∈ A -> rtc HRed.R a b -> Γ ⊢ a ≡ b ∈ A.
+  Lemma ToEq Γ (a b : PTm) A : Γ ⊢ a ∈ A -> rtc HRed.R a b -> Γ ⊢ a ≡ b ∈ A.
   Proof.
     induction 2; sauto lq:on use:HRed.ToEq, E_Transitive, HRed.preservation.
   Qed.
 End HReds.
 
-Lemma T_Conv_E n Γ (a : PTm n) A B i :
+Lemma T_Conv_E Γ (a : PTm) A B i :
   Γ ⊢ a ∈ A ->
   Γ ⊢ A ≡ B ∈ PUniv i \/ Γ ⊢ B ≡ A ∈ PUniv i ->
   Γ ⊢ a ∈ B.
 Proof. qauto use:T_Conv, Su_Eq, E_Symmetric. Qed.
 
-Lemma E_Conv_E n Γ (a b : PTm n) A B i :
+Lemma E_Conv_E Γ (a b : PTm) A B i :
   Γ ⊢ a ≡ b ∈ A ->
   Γ ⊢ A ≡ B ∈ PUniv i \/ Γ ⊢ B ≡ A ∈ PUniv i ->
   Γ ⊢ a ≡ b ∈ B.
 Proof. qauto use:E_Conv, Su_Eq, E_Symmetric. Qed.
 
-Lemma lored_embed n Γ (a b : PTm n) A :
+Lemma lored_embed Γ (a b : PTm) A :
   Γ ⊢ a ∈ A -> LoRed.R a b -> Γ ⊢ a ≡ b ∈ A.
 Proof. sfirstorder use:LoRed.ToRRed, RRed_Eq. Qed.
 
-Lemma loreds_embed n Γ (a b : PTm n) A :
+Lemma loreds_embed Γ (a b : PTm ) A :
   Γ ⊢ a ∈ A -> rtc LoRed.R a b -> Γ ⊢ a ≡ b ∈ A.
 Proof.
   move => + h. move : Γ A.
@@ -84,25 +84,18 @@ Proof.
     hauto lq:on ctrs:Eq.
 Qed.
 
-Lemma T_Bot_Imp n Γ (A : PTm n) :
-  Γ ⊢ PBot ∈ A -> False.
-  move E : PBot => u hu.
-  move : E.
-  induction hu => //=.
-Qed.
-
-Lemma Zero_Inv n Γ U :
-  Γ ⊢ @PZero n ∈ U ->
+Lemma Zero_Inv Γ U :
+  Γ ⊢ PZero ∈ U ->
   Γ ⊢ PNat ≲ U.
 Proof.
   move E : PZero => u hu.
   move : E.
-  elim : n Γ u U /hu=>//=.
+  elim : Γ u U /hu=>//=.
   by eauto using Su_Eq, E_Refl, T_Nat'.
   hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma Sub_Bind_InvR n Γ p (A : PTm n) B C :
+Lemma Sub_Bind_InvR Γ p (A : PTm ) B C :
   Γ ⊢ PBind p A B ≲ C ->
   exists i A0 B0, Γ ⊢ C ≡ PBind p A0 B0 ∈ PUniv i.
 Proof.
@@ -140,7 +133,6 @@ Proof.
     have ? : b = p by hauto l:on use:Sub.bind_inj. subst.
     eauto.
   - hauto lq:on use:synsub_to_usub, Sub.bind_univ_noconf.
-  - hauto lq:on use:regularity, T_Bot_Imp.
   - move => _ _ /synsub_to_usub [_ [_ h]]. exfalso.
     apply Sub.nat_bind_noconf in h => //=.
   - move => h.
@@ -158,7 +150,7 @@ Proof.
     exfalso. move : h2 hne. hauto l:on use:Sub.bind_sne_noconf.
 Qed.
 
-Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
+Lemma Sub_Univ_InvR Γ i C :
   Γ ⊢ PUniv i ≲ C ->
   exists j k, Γ ⊢ C ≡ PUniv j ∈ PUniv k.
 Proof.
@@ -188,7 +180,6 @@ Proof.
   - hauto q:on use:synsub_to_usub, Sub.univ_sne_noconf, ne_nf_embed.
   - hauto lq:on use:synsub_to_usub, Sub.univ_bind_noconf.
   - sfirstorder.
-  - hauto lq:on use:regularity, T_Bot_Imp.
   - hauto q:on use:synsub_to_usub, Sub.nat_univ_noconf.
   - move => h.
     have {}h : Γ ⊢ PZero ∈ PUniv j by hauto l:on use:regularity.
@@ -205,7 +196,7 @@ Proof.
     exfalso. move : h2 hne. hauto l:on use:Sub.univ_sne_noconf.
 Qed.
 
-Lemma Sub_Bind_InvL n Γ p (A : PTm n) B C :
+Lemma Sub_Bind_InvL Γ p (A : PTm) B C :
   Γ ⊢ C ≲ PBind p A B ->
   exists i A0 B0, Γ ⊢ PBind p A0 B0 ≡ C ∈ PUniv i.
 Proof.
@@ -243,7 +234,6 @@ Proof.
     have ? : b = p by hauto l:on use:Sub.bind_inj. subst.
     eauto using E_Symmetric.
   - hauto lq:on use:synsub_to_usub, Sub.univ_bind_noconf.
-  - hauto lq:on use:regularity, T_Bot_Imp.
   - move => _ _ /synsub_to_usub [_ [_ h]]. exfalso.
     apply Sub.bind_nat_noconf in h => //=.
   - move => h.
@@ -262,7 +252,7 @@ Proof.
     hauto l:on use:Sub.sne_bind_noconf.
 Qed.
 
-Lemma T_Abs_Inv n Γ (a0 a1 : PTm (S n)) U :
+Lemma T_Abs_Inv Γ (a0 a1 : PTm) U :
   Γ ⊢ PAbs a0 ∈ U ->
   Γ ⊢ PAbs a1 ∈ U ->
   exists Δ V, Δ ⊢ a0 ∈ V /\ Δ ⊢ a1 ∈ V.
@@ -272,7 +262,7 @@ Proof.
   move /Sub_Bind_InvR : (hU0) => [i][A2][B2]hE.
   have hSu : Γ ⊢ PBind PPi A1 B1 ≲ PBind PPi A2 B2 by eauto using Su_Eq, Su_Transitive.
   have hSu' : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Eq, Su_Transitive.
-  exists (funcomp (ren_PTm shift) (scons A2 Γ)), B2.
+  exists ((cons A2 Γ)), B2.
   have [k ?] : exists k, Γ ⊢ A2 ∈ PUniv k by hauto lq:on use:regularity, Bind_Inv.
   split.
   - have /Su_Pi_Proj2_Var ? := hSu'.
@@ -287,15 +277,15 @@ Proof.
     apply : ctx_eq_subst_one; eauto.
 Qed.
 
-Lemma T_Univ_Raise n Γ (a : PTm n) i j :
+Lemma T_Univ_Raise Γ (a : PTm) i j :
   Γ ⊢ a ∈ PUniv i ->
   i <= j ->
   Γ ⊢ a ∈ PUniv j.
 Proof. hauto lq:on rew:off use:T_Conv, Su_Univ, wff_mutual. Qed.
 
-Lemma Bind_Univ_Inv n Γ p (A : PTm n) B i :
+Lemma Bind_Univ_Inv Γ p (A : PTm) B i :
   Γ ⊢ PBind p A B ∈ PUniv i ->
-  Γ ⊢ A ∈ PUniv i /\ funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i.
+  Γ ⊢ A ∈ PUniv i /\ (cons A Γ) ⊢ B ∈ PUniv i.
 Proof.
   move /Bind_Inv.
   move => [i0][hA][hB]h.
@@ -303,9 +293,9 @@ Proof.
   sfirstorder use:T_Univ_Raise.
 Qed.
 
-Lemma Abs_Pi_Inv n Γ (a : PTm (S n)) A B :
+Lemma Abs_Pi_Inv Γ (a : PTm) A B :
   Γ ⊢ PAbs a ∈ PBind PPi A B ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B.
+  (cons A Γ) ⊢ a ∈ B.
 Proof.
   move => h.
   have [i hi] : exists i, Γ ⊢ PBind PPi A B ∈ PUniv i by hauto use:regularity.
@@ -314,15 +304,16 @@ Proof.
   apply : T_App'; cycle 1.
   apply : weakening_wt'; cycle 2. apply hi.
   apply h. reflexivity. reflexivity. rewrite -/ren_PTm.
-  apply T_Var' with (i := var_zero).  by asimpl.
+  apply T_Var with (i := var_zero).
   by eauto using Wff_Cons'.
+  apply here.
   rewrite -/ren_PTm.
-  by asimpl.
+  by asimpl; rewrite subst_scons_id.
   rewrite -/ren_PTm.
-  by asimpl.
+  by asimpl; rewrite subst_scons_id.
 Qed.
 
-Lemma Pair_Sig_Inv n Γ (a b : PTm n) A B :
+Lemma Pair_Sig_Inv Γ (a b : PTm) A B :
   Γ ⊢ PPair a b ∈ PBind PSig A B ->
   Γ ⊢ a ∈ A /\ Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B.
 Proof.
@@ -345,7 +336,7 @@ Qed.
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
 Reserved Notation "a ⇔ b" (at level 70).
-Inductive CoqEq {n} : PTm n -> PTm n -> Prop :=
+Inductive CoqEq : PTm -> PTm -> Prop :=
 | CE_ZeroZero :
   PZero ↔ PZero
 
@@ -409,7 +400,7 @@ Inductive CoqEq {n} : PTm n -> PTm n -> Prop :=
   a0 ∼ a1 ->
   a0 ↔ a1
 
-with CoqEq_Neu  {n} : PTm n -> PTm n -> Prop :=
+with CoqEq_Neu : PTm -> PTm -> Prop :=
 | CE_VarCong i :
   (* -------------------------- *)
   VarPTm i ∼ VarPTm i
@@ -439,7 +430,7 @@ with CoqEq_Neu  {n} : PTm n -> PTm n -> Prop :=
   (* ----------------------------------- *)
   PInd P0 u0 b0 c0 ∼ PInd P1 u1 b1 c1
 
-with CoqEq_R {n} : PTm n -> PTm n -> Prop :=
+with CoqEq_R : PTm -> PTm -> Prop :=
 | CE_HRed a a' b b'  :
   rtc HRed.R a a' ->
   rtc HRed.R b b' ->
@@ -448,7 +439,7 @@ with CoqEq_R {n} : PTm n -> PTm n -> Prop :=
   a ⇔ b
 where "a ↔ b" := (CoqEq a b) and "a ⇔ b" := (CoqEq_R a b) and "a ∼ b" := (CoqEq_Neu a b).
 
-Lemma CE_HRedL n (a a' b : PTm n)  :
+Lemma CE_HRedL (a a' b : PTm)  :
   HRed.R a a' ->
   a' ⇔ b ->
   a ⇔ b.
@@ -463,27 +454,28 @@ Scheme
 
 Combined Scheme coqeq_mutual from coqeq_neu_ind, coqeq_ind, coqeq_r_ind.
 
-Lemma coqeq_symmetric_mutual : forall n,
-    (forall (a b : PTm n), a ∼ b -> b ∼ a) /\
-    (forall (a b : PTm n), a ↔ b -> b ↔ a) /\
-    (forall (a b : PTm n), a ⇔ b -> b ⇔ a).
+Lemma coqeq_symmetric_mutual :
+    (forall (a b : PTm), a ∼ b -> b ∼ a) /\
+    (forall (a b : PTm), a ↔ b -> b ↔ a) /\
+    (forall (a b : PTm), a ⇔ b -> b ⇔ a).
 Proof. apply coqeq_mutual; qauto l:on ctrs:CoqEq,CoqEq_R, CoqEq_Neu. Qed.
 
-Lemma coqeq_sound_mutual : forall n,
-    (forall (a b : PTm n), a ∼ b -> forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> exists C,
+Lemma coqeq_sound_mutual :
+    (forall (a b : PTm ), a ∼ b -> forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> exists C,
        Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ≡ b ∈ C) /\
-    (forall (a b : PTm n), a ↔ b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> Γ ⊢ a ≡ b ∈ A) /\
-    (forall (a b : PTm n), a ⇔ b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> Γ ⊢ a ≡ b ∈ A).
+    (forall (a b : PTm ), a ↔ b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> Γ ⊢ a ≡ b ∈ A) /\
+    (forall (a b : PTm ), a ⇔ b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> Γ ⊢ a ≡ b ∈ A).
 Proof.
   move => [:hAppL hPairL].
   apply coqeq_mutual.
-  - move => n i Γ A B hi0 hi1.
-    move /Var_Inv : hi0 => [hΓ h0].
-    move /Var_Inv : hi1 => [_ h1].
-    exists (Γ i).
+  - move => i Γ A B hi0 hi1.
+    move /Var_Inv : hi0 => [hΓ [C [h00 h01]]].
+    move /Var_Inv : hi1 => [_ [C0 [h10 h11]]].
+    have ? : C0 = C by eauto using lookup_deter. subst.
+    exists C.
     repeat split => //=.
     apply E_Refl. eauto using T_Var.
-  - move => n [] u0 u1 hu0 hu1 hu ihu Γ A B hu0' hu1'.
+  - move => [] u0 u1 hu0 hu1 hu ihu Γ A B hu0' hu1'.
     + move /Proj1_Inv : hu0'.
       move => [A0][B0][hu0']hu0''.
       move /Proj1_Inv : hu1'.
@@ -517,7 +509,7 @@ Proof.
       apply : Su_Sig_Proj2; eauto.
       apply : E_Proj1; eauto.
       apply : E_Proj2; eauto.
-  - move => n u0 u1 a0 a1 neu0 neu1 hu ihu ha iha Γ A B wta0 wta1.
+  - move => u0 u1 a0 a1 neu0 neu1 hu ihu ha iha Γ A B wta0 wta1.
     move /App_Inv : wta0 => [A0][B0][hu0][ha0]hU.
     move /App_Inv : wta1 => [A1][B1][hu1][ha1]hU1.
     move : ihu hu0 hu1. repeat move/[apply].
@@ -539,13 +531,13 @@ Proof.
       first by sfirstorder use:bind_inst.
     apply : Su_Pi_Proj2';  eauto using E_Refl.
     apply E_App with (A := A2); eauto using E_Conv_E.
-  - move {hAppL hPairL} => n P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 hP ihP hu ihu hb ihb hc ihc Γ A B.
+  - move {hAppL hPairL} => P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 hP ihP hu ihu hb ihb hc ihc Γ A B.
     move /Ind_Inv => [i0][hP0][hu0][hb0][hc0]hSu0.
     move /Ind_Inv => [i1][hP1][hu1][hb1][hc1]hSu1.
     move : ihu hu0 hu1; do!move/[apply]. move => ihu.
     have {}ihu : Γ ⊢ u0 ≡ u1 ∈ PNat by hauto l:on use:E_Conv.
     have wfΓ : ⊢ Γ by hauto use:wff_mutual.
-    have wfΓ' : ⊢ funcomp (ren_PTm shift) (scons PNat Γ) by hauto lq:on use:Wff_Cons', T_Nat'.
+    have wfΓ' : ⊢ (cons PNat Γ) by hauto lq:on use:Wff_Cons', T_Nat'.
     move => [:sigeq].
     have sigeq' : Γ ⊢ PBind PSig PNat P0 ≡ PBind PSig PNat P1 ∈ PUniv (max i0 i1).
     apply E_Bind. by eauto using T_Nat, E_Refl.
@@ -567,17 +559,17 @@ Proof.
     apply morphing_ext. set x := {1}(funcomp _ shift).
     have -> : x = funcomp (ren_PTm shift) VarPTm by asimpl.
     apply : morphing_ren; eauto. apply : renaming_shift; eauto. by apply morphing_id.
-    apply T_Suc. by apply T_Var.
+    apply T_Suc. apply T_Var; eauto using here.
   - hauto lq:on use:Zero_Inv db:wt.
   - hauto lq:on use:Suc_Inv db:wt.
-  - move => n a b ha iha Γ A h0 h1.
+  - move => a b ha iha Γ A h0 h1.
     move /Abs_Inv : h0 => [A0][B0][h0]h0'.
     move /Abs_Inv : h1 => [A1][B1][h1]h1'.
     have [i [A2 [B2 h]]] : exists i A2 B2, Γ ⊢ A ≡ PBind PPi A2 B2 ∈ PUniv i by hauto l:on use:Sub_Bind_InvR.
     have hp0 : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
     have hp1 : Γ ⊢ PBind PPi A1 B1 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
-    have [j ?] : exists j, Γ ⊢ A0 ∈ PUniv j by qauto l:on use:Wff_Cons_Inv, wff_mutual.
-    have [k ?] : exists j, Γ ⊢ A1 ∈ PUniv j by qauto l:on use:Wff_Cons_Inv, wff_mutual.
+    have [j ?] : exists j, Γ ⊢ A0 ∈ PUniv j by eapply wff_mutual in h0; inversion h0; subst; eauto.
+    have [k ?] : exists j, Γ ⊢ A1 ∈ PUniv j by eapply wff_mutual in h1; inversion h1; subst; eauto.
     have [l ?] : exists j, Γ ⊢ A2 ∈ PUniv j by hauto lq:on rew:off use:regularity, Bind_Inv.
     have [h2 h3] : Γ ⊢ A2 ≲ A0 /\ Γ ⊢ A2 ≲ A1 by hauto l:on use:Su_Pi_Proj1.
     apply E_Conv with (A := PBind PPi A2 B2); cycle 1.
@@ -591,7 +583,7 @@ Proof.
     apply : T_Conv; eauto.
     eapply ctx_eq_subst_one with (A0 := A1); eauto.
   - abstract : hAppL.
-    move => n a u hneu ha iha Γ A wta wtu.
+    move => a u hneu ha iha Γ A wta wtu.
     move /Abs_Inv : wta => [A0][B0][wta]hPi.
     have [i [A2 [B2 h]]] : exists i A2 B2, Γ ⊢ A ≡ PBind PPi A2 B2 ∈ PUniv i by hauto l:on use:Sub_Bind_InvR.
     have hPi'' : Γ ⊢ PBind PPi A2 B2 ≲ A by eauto using Su_Eq, Su_Transitive, E_Symmetric.
@@ -603,7 +595,6 @@ Proof.
     have /regularity_sub0 [i' hPi0] := hPi.
     have : Γ ⊢ PAbs (PApp (ren_PTm shift u) (VarPTm var_zero)) ≡ u ∈ PBind PPi A2 B2.
     apply : E_AppEta; eauto.
-    sfirstorder use:wff_mutual.
     hauto l:on use:regularity.
     apply T_Conv with (A := A);eauto.
     eauto using Su_Eq.
@@ -616,19 +607,18 @@ Proof.
     apply : T_Conv; eauto.
     eapply ctx_eq_subst_one with (A0 := A0); eauto.
     sfirstorder use:Su_Pi_Proj1.
-
-    (* move /Su_Pi_Proj2_Var in hPidup. *)
-    (* apply : T_Conv; eauto. *)
-    eapply T_App' with (A := ren_PTm shift A2) (B := ren_PTm (upRen_PTm_PTm shift) B2). by asimpl.
+    eapply T_App' with (A := ren_PTm shift A2) (B := ren_PTm (upRen_PTm_PTm shift) B2).
+    by asimpl; rewrite subst_scons_id.
     eapply weakening_wt' with (a := u) (A := PBind PPi A2 B2);eauto.
     by eauto using T_Conv_E. apply T_Var. apply : Wff_Cons'; eauto.
+    apply here.
   (* Mirrors the last case *)
-  - move => n a u hu ha iha Γ A hu0 ha0.
+  - move => a u hu ha iha Γ A hu0 ha0.
     apply E_Symmetric.
     apply : hAppL; eauto.
     sfirstorder use:coqeq_symmetric_mutual.
     sfirstorder use:E_Symmetric.
-  - move => {hAppL hPairL} n a0 a1 b0 b1 ha iha hb ihb Γ A.
+  - move => {hAppL hPairL} a0 a1 b0 b1 ha iha hb ihb Γ A.
     move /Pair_Inv => [A0][B0][h00][h01]h02.
     move /Pair_Inv => [A1][B1][h10][h11]h12.
     have [i[A2[B2 h2]]] : exists i A2 B2, Γ ⊢ A ≡ PBind PSig A2 B2 ∈ PUniv i by hauto l:on use:Sub_Bind_InvR.
@@ -654,7 +644,7 @@ Proof.
   - move => {hAppL}.
     abstract : hPairL.
     move => {hAppL}.
-    move => n a0 a1 u neu h0 ih0 h1 ih1 Γ A ha hu.
+    move => a0 a1 u neu h0 ih0 h1 ih1 Γ A ha hu.
     move /Pair_Inv : ha => [A0][B0][ha0][ha1]ha.
     have [i [A2 [B2 hA]]] : exists i A2 B2, Γ ⊢ A ≡ PBind PSig A2 B2 ∈ PUniv i by hauto l:on use:Sub_Bind_InvR.
     have hA' : Γ ⊢ PBind PSig A2 B2 ≲ A by eauto using E_Symmetric, Su_Eq.
@@ -683,7 +673,7 @@ Proof.
     move => *. apply E_Symmetric. apply : hPairL;
       sfirstorder use:coqeq_symmetric_mutual, E_Symmetric.
   - sfirstorder use:E_Refl.
-  - move => {hAppL hPairL} n p A0 A1 B0 B1 hA ihA hB ihB Γ A hA0 hA1.
+  - move => {hAppL hPairL} p A0 A1 B0 B1 hA ihA hB ihB Γ A hA0 hA1.
     move /Bind_Inv : hA0 => [i][hA0][hB0]hU.
     move /Bind_Inv : hA1 => [j][hA1][hB1]hU1.
     have [l [k hk]] : exists l k, Γ ⊢ A ≡ PUniv k ∈ PUniv l
@@ -703,23 +693,23 @@ Proof.
     move : weakening_su hjk hA0. by repeat move/[apply].
   - hauto lq:on ctrs:Eq,LEq,Wt.
   - hauto lq:on ctrs:Eq,LEq,Wt.
-  - move => n a a' b b' ha hb hab ihab Γ A ha0 hb0.
+  - move => a a' b b' ha hb hab ihab Γ A ha0 hb0.
     have [*] : Γ ⊢ a' ∈ A /\ Γ ⊢ b' ∈ A by eauto using HReds.preservation.
     hauto lq:on use:HReds.ToEq, E_Symmetric, E_Transitive.
 Qed.
 
-Definition algo_metric {n} k (a b : PTm n) :=
+Definition algo_metric k (a b : PTm) :=
   exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ nf va /\ nf vb /\ EJoin.R va vb /\ size_PTm va + size_PTm vb + i + j <= k.
 
-Lemma ne_hne n (a : PTm n) : ne a -> ishne a.
+Lemma ne_hne (a : PTm) : ne a -> ishne a.
 Proof. elim : a => //=; sfirstorder b:on. Qed.
 
-Lemma hf_hred_lored n (a b : PTm n) :
+Lemma hf_hred_lored (a b : PTm) :
   ~~ ishf a ->
   LoRed.R a b ->
   HRed.R a b \/ ishne a.
 Proof.
-  move => + h. elim : n a b/ h=>//=.
+  move => + h. elim : a b/ h=>//=.
   - hauto l:on use:HRed.AppAbs.
   - hauto l:on use:HRed.ProjPair.
   - hauto lq:on ctrs:HRed.R.
@@ -733,7 +723,7 @@ Proof.
   - hauto lq:on use:ne_hne.
 Qed.
 
-Lemma algo_metric_case n k (a b : PTm n) :
+Lemma algo_metric_case k (a b : PTm) :
   algo_metric k a b ->
   (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ algo_metric k' a' b /\ k' < k.
 Proof.
@@ -761,28 +751,21 @@ Proof.
     + sfirstorder use:ne_hne.
 Qed.
 
-Lemma algo_metric_sym n k (a b : PTm n) :
+Lemma algo_metric_sym k (a b : PTm) :
   algo_metric k a b -> algo_metric k b a.
 Proof. hauto lq:on unfold:algo_metric solve+:lia. Qed.
 
-Lemma hred_hne n (a b : PTm n) :
+Lemma hred_hne (a b : PTm) :
   HRed.R a b ->
   ishne a ->
   False.
 Proof. induction 1; sfirstorder. Qed.
 
-Lemma hf_not_hne n (a : PTm n) :
+Lemma hf_not_hne (a : PTm) :
   ishf a -> ishne a -> False.
 Proof. case : a => //=. Qed.
 
-(* Lemma algo_metric_hf_case n Γ k a b (A : PTm n) : *)
-(*   Γ ⊢ a ∈ A -> *)
-(*   Γ ⊢ b ∈ A -> *)
-(*   algo_metric k a b -> ishf a -> ishf b -> *)
-(*   (exists a' b' k', a = PAbs a' /\ b = PAbs b' /\ algo_metric k' a' b' /\ k' < k) \/ *)
-(*   (exists a0' a1' b0' b1' ka kb, a = PPair a0' a1' /\ b = PPair b0' b1' /\ algo_metric ka a0' b0' /\ algo_metric ) *)
-
-Lemma T_AbsPair_Imp n Γ a (b0 b1 : PTm n) A :
+Lemma T_AbsPair_Imp Γ a (b0 b1 : PTm) A :
   Γ ⊢ PAbs a ∈ A ->
   Γ ⊢ PPair b0 b1 ∈ A ->
   False.
@@ -794,7 +777,7 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_inj.
 Qed.
 
-Lemma T_AbsZero_Imp n Γ a (A : PTm n) :
+Lemma T_AbsZero_Imp Γ a (A : PTm) :
   Γ ⊢ PAbs a ∈ A ->
   Γ ⊢ PZero ∈ A ->
   False.
@@ -806,7 +789,7 @@ Proof.
   clear. hauto lq:on use:synsub_to_usub, Sub.bind_nat_noconf.
 Qed.
 
-Lemma T_AbsSuc_Imp n Γ a b (A : PTm n) :
+Lemma T_AbsSuc_Imp Γ a b (A : PTm) :
   Γ ⊢ PAbs a ∈ A ->
   Γ ⊢ PSuc b ∈ A ->
   False.
@@ -818,18 +801,18 @@ Proof.
   hauto lq:on use:Sub.bind_nat_noconf, synsub_to_usub.
 Qed.
 
-Lemma Nat_Inv n Γ A:
-  Γ ⊢ @PNat n ∈ A ->
+Lemma Nat_Inv Γ A:
+  Γ ⊢ PNat ∈ A ->
   exists i, Γ ⊢ PUniv i ≲ A.
 Proof.
   move E : PNat => u hu.
   move : E.
-  elim : n Γ u A / hu=>//=.
+  elim : Γ u A / hu=>//=.
   - hauto lq:on use:E_Refl, T_Univ, Su_Eq.
   - hauto lq:on ctrs:LEq.
 Qed.
 
-Lemma T_AbsNat_Imp n Γ a (A : PTm n) :
+Lemma T_AbsNat_Imp Γ a (A : PTm ) :
   Γ ⊢ PAbs a ∈ A ->
   Γ ⊢ PNat ∈ A ->
   False.
@@ -841,7 +824,7 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf, synsub_to_usub.
 Qed.
 
-Lemma T_PairBind_Imp n Γ (a0 a1 : PTm n) p A0 B0 U :
+Lemma T_PairBind_Imp Γ (a0 a1 : PTm ) p A0 B0 U :
   Γ ⊢ PPair a0 a1 ∈ U ->
   Γ ⊢ PBind p A0 B0  ∈ U ->
   False.
@@ -853,7 +836,7 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_univ_noconf.
 Qed.
 
-Lemma T_PairNat_Imp n Γ (a0 a1 : PTm n) U :
+Lemma T_PairNat_Imp Γ (a0 a1 : PTm) U :
   Γ ⊢ PPair a0 a1 ∈ U ->
   Γ ⊢ PNat ∈ U ->
   False.
@@ -864,7 +847,7 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_univ_noconf.
 Qed.
 
-Lemma T_PairZero_Imp n Γ (a0 a1 : PTm n) U :
+Lemma T_PairZero_Imp Γ (a0 a1 : PTm ) U :
   Γ ⊢ PPair a0 a1 ∈ U ->
   Γ ⊢ PZero ∈ U ->
   False.
@@ -875,7 +858,7 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_nat_noconf.
 Qed.
 
-Lemma T_PairSuc_Imp n Γ (a0 a1 : PTm n) b U :
+Lemma T_PairSuc_Imp Γ (a0 a1 : PTm ) b U :
   Γ ⊢ PPair a0 a1 ∈ U ->
   Γ ⊢ PSuc b ∈ U ->
   False.
@@ -886,17 +869,17 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_nat_noconf.
 Qed.
 
-Lemma Univ_Inv n Γ i U :
-  Γ ⊢ @PUniv n i ∈ U ->
+Lemma Univ_Inv Γ i U :
+  Γ ⊢ PUniv i ∈ U ->
   Γ ⊢ PUniv i ∈ PUniv (S i)  /\ Γ ⊢ PUniv (S i) ≲ U.
 Proof.
   move E : (PUniv i) => u hu.
-  move : i E. elim : n Γ u U / hu => n //=.
+  move : i E. elim : Γ u U / hu => n //=.
   - hauto l:on use:E_Refl, Su_Eq, T_Univ.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma T_PairUniv_Imp n Γ (a0 a1 : PTm n) i U :
+Lemma T_PairUniv_Imp Γ (a0 a1 : PTm) i U :
   Γ ⊢ PPair a0 a1 ∈ U ->
   Γ ⊢ PUniv i ∈ U ->
   False.
@@ -909,7 +892,7 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
-Lemma T_AbsUniv_Imp n Γ a i (A : PTm n) :
+Lemma T_AbsUniv_Imp Γ a i (A : PTm) :
   Γ ⊢ PAbs a ∈ A ->
   Γ ⊢ PUniv i ∈ A ->
   False.
@@ -922,19 +905,19 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
-Lemma T_AbsUniv_Imp' n Γ (a : PTm (S n)) i  :
+Lemma T_AbsUniv_Imp' Γ (a : PTm) i  :
   Γ ⊢ PAbs a ∈ PUniv i -> False.
 Proof.
   hauto lq:on use:synsub_to_usub, Sub.bind_univ_noconf, Abs_Inv.
 Qed.
 
-Lemma T_PairUniv_Imp' n Γ (a b : PTm n) i  :
+Lemma T_PairUniv_Imp' Γ (a b : PTm) i  :
   Γ ⊢ PPair a b ∈ PUniv i -> False.
 Proof.
   hauto lq:on use:synsub_to_usub, Sub.bind_univ_noconf, Pair_Inv.
 Qed.
 
-Lemma T_AbsBind_Imp n Γ a p A0 B0 (U : PTm n) :
+Lemma T_AbsBind_Imp Γ a p A0 B0 (U : PTm ) :
   Γ ⊢ PAbs a ∈ U ->
   Γ ⊢ PBind p A0 B0 ∈ U ->
   False.
@@ -947,7 +930,7 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
-Lemma lored_nsteps_suc_inv k n (a : PTm n) b :
+Lemma lored_nsteps_suc_inv k (a : PTm ) b :
   nsteps LoRed.R k (PSuc a) b -> exists b', nsteps LoRed.R k a b' /\ b = PSuc b'.
 Proof.
   move E : (PSuc a) => u hu.
@@ -957,7 +940,7 @@ Proof.
   - scrush ctrs:nsteps inv:LoRed.R.
 Qed.
 
-Lemma lored_nsteps_abs_inv k n (a : PTm (S n)) b :
+Lemma lored_nsteps_abs_inv k (a : PTm) b :
   nsteps LoRed.R k (PAbs a) b -> exists b', nsteps LoRed.R k a b' /\ b = PAbs b'.
 Proof.
   move E : (PAbs a) => u hu.
@@ -967,15 +950,15 @@ Proof.
   - scrush ctrs:nsteps inv:LoRed.R.
 Qed.
 
-Lemma lored_hne_preservation n (a b : PTm n) :
+Lemma lored_hne_preservation (a b : PTm) :
     LoRed.R a b -> ishne a -> ishne b.
 Proof.  induction 1; sfirstorder. Qed.
 
-Lemma loreds_hne_preservation n (a b : PTm n) :
+Lemma loreds_hne_preservation (a b : PTm ) :
   rtc LoRed.R a b -> ishne a -> ishne b.
 Proof. induction 1; hauto lq:on ctrs:rtc use:lored_hne_preservation. Qed.
 
-Lemma lored_nsteps_bind_inv k n p (a0 : PTm n) b0 C :
+Lemma lored_nsteps_bind_inv k p (a0 : PTm ) b0 C :
     nsteps LoRed.R k (PBind p a0 b0) C ->
     exists i j a1 b1,
       i <= k /\ j <= k /\
@@ -995,7 +978,7 @@ Proof.
     exists i,(S j),a1,b1. sauto lq:on solve+:lia.
 Qed.
 
-Lemma lored_nsteps_ind_inv k n P0 (a0 : PTm n) b0 c0 U :
+Lemma lored_nsteps_ind_inv k P0 (a0 : PTm) b0 c0 U :
   nsteps LoRed.R k (PInd P0 a0 b0 c0) U ->
   ishne a0 ->
   exists iP ia ib ic P1 a1 b1 c1,
@@ -1026,7 +1009,7 @@ Proof.
       exists iP, ia, ib, (S ic). sauto lq:on solve+:lia.
 Qed.
 
-Lemma lored_nsteps_app_inv k n (a0 b0 C : PTm n) :
+Lemma lored_nsteps_app_inv k (a0 b0 C : PTm) :
     nsteps LoRed.R k (PApp a0 b0) C ->
     ishne a0 ->
     exists i j a1 b1,
@@ -1050,7 +1033,7 @@ Proof.
     exists i, (S j), a1, b1. sauto lq:on solve+:lia.
 Qed.
 
-Lemma lored_nsteps_proj_inv k n p (a0 C : PTm n) :
+Lemma lored_nsteps_proj_inv k p (a0 C : PTm) :
     nsteps LoRed.R k (PProj p a0) C ->
     ishne a0 ->
     exists i a1,
@@ -1070,7 +1053,7 @@ Proof.
     exists (S i), a1. hauto lq:on ctrs:nsteps solve+:lia.
 Qed.
 
-Lemma algo_metric_proj n k p0 p1 (a0 a1 : PTm n) :
+Lemma algo_metric_proj k p0 p1 (a0 a1 : PTm) :
   algo_metric k (PProj p0 a0) (PProj p1 a1) ->
   ishne a0 ->
   ishne a1 ->
@@ -1092,7 +1075,7 @@ Proof.
 Qed.
 
 
-Lemma lored_nsteps_app_cong k n (a0 a1 b : PTm n) :
+Lemma lored_nsteps_app_cong k (a0 a1 b : PTm) :
   nsteps LoRed.R k a0 a1 ->
   ishne a0 ->
   nsteps LoRed.R k (PApp a0 b) (PApp a1 b).
@@ -1106,7 +1089,7 @@ Lemma lored_nsteps_app_cong k n (a0 a1 b : PTm n) :
     apply ih. sfirstorder use:lored_hne_preservation.
 Qed.
 
-Lemma lored_nsteps_proj_cong k n p (a0 a1 : PTm n) :
+Lemma lored_nsteps_proj_cong k p (a0 a1 : PTm) :
   nsteps LoRed.R k a0 a1 ->
   ishne a0 ->
   nsteps LoRed.R k (PProj p a0) (PProj p a1).
@@ -1119,7 +1102,7 @@ Lemma lored_nsteps_proj_cong k n p (a0 a1 : PTm n) :
     apply ih. sfirstorder use:lored_hne_preservation.
 Qed.
 
-Lemma lored_nsteps_pair_inv k n (a0 b0 C : PTm n) :
+Lemma lored_nsteps_pair_inv k (a0 b0 C : PTm ) :
     nsteps LoRed.R k (PPair a0 b0) C ->
     exists i j a1 b1,
       i <= k /\ j <= k /\
@@ -1139,7 +1122,7 @@ Lemma lored_nsteps_pair_inv k n (a0 b0 C : PTm n) :
     exists i, (S j), a1, b1. sauto lq:on solve+:lia.
 Qed.
 
-Lemma algo_metric_join n k (a b : PTm n) :
+Lemma algo_metric_join k (a b : PTm ) :
   algo_metric k a b ->
   DJoin.R a b.
   rewrite /algo_metric.
@@ -1152,7 +1135,7 @@ Lemma algo_metric_join n k (a b : PTm n) :
   rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
 Qed.
 
-Lemma algo_metric_pair n k (a0 b0 a1 b1 : PTm n) :
+Lemma algo_metric_pair k (a0 b0 a1 b1 : PTm) :
   SN (PPair a0 b0) ->
   SN (PPair a1 b1) ->
   algo_metric k (PPair a0 b0) (PPair a1 b1) ->
@@ -1170,18 +1153,18 @@ Lemma algo_metric_pair n k (a0 b0 a1 b1 : PTm n) :
   hauto l:on use:DJoin.ejoin_pair_inj.
 Qed.
 
-Lemma hne_nf_ne n (a : PTm n) :
+Lemma hne_nf_ne (a : PTm ) :
   ishne a -> nf a -> ne a.
 Proof. case : a => //=. Qed.
 
-Lemma lored_nsteps_renaming k n m (a b : PTm n) (ξ : fin n -> fin m) :
+Lemma lored_nsteps_renaming k (a b : PTm) (ξ : nat -> nat) :
   nsteps LoRed.R k a b ->
   nsteps LoRed.R k (ren_PTm ξ a) (ren_PTm ξ b).
 Proof.
   induction 1; hauto lq:on ctrs:nsteps use:LoRed.renaming.
 Qed.
 
-Lemma algo_metric_neu_abs n k (a0 : PTm (S n)) u :
+Lemma algo_metric_neu_abs k (a0 : PTm) u :
   algo_metric k u (PAbs a0) ->
   ishne u ->
   exists j, j < k /\ algo_metric j (PApp (ren_PTm shift u) (VarPTm var_zero)) a0.
@@ -1203,7 +1186,7 @@ Proof.
   simpl in *.  rewrite size_PTm_ren. lia.
 Qed.
 
-Lemma algo_metric_neu_pair n k (a0 b0 : PTm n) u :
+Lemma algo_metric_neu_pair k (a0 b0 : PTm) u :
   algo_metric k u (PPair a0 b0) ->
   ishne u ->
   exists j, j < k /\ algo_metric j (PProj PL u) a0 /\ algo_metric j (PProj PR u) b0.
@@ -1224,7 +1207,7 @@ Proof.
   eapply DJoin.ejoin_pair_inj; hauto qb:on ctrs:rtc, ERed.R.
 Qed.
 
-Lemma algo_metric_ind n k P0 (a0 : PTm n) b0 c0 P1 a1 b1 c1 :
+Lemma algo_metric_ind k P0 (a0 : PTm ) b0 c0 P1 a1 b1 c1 :
   algo_metric k (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
   ishne a0 ->
   ishne a1 ->
@@ -1242,7 +1225,7 @@ Proof.
   hauto lq:on rew:off use:ne_nf b:on solve+:lia.
 Qed.
 
-Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
+Lemma algo_metric_app k (a0 b0 a1 b1 : PTm) :
   algo_metric k (PApp a0 b0) (PApp a1 b1) ->
   ishne a0 ->
   ishne a1 ->
@@ -1271,7 +1254,7 @@ Proof.
     repeat split => //=; sfirstorder b:on use:ne_nf.
 Qed.
 
-Lemma algo_metric_suc n k (a0 a1 : PTm n) :
+Lemma algo_metric_suc k (a0 a1 : PTm) :
   algo_metric k (PSuc a0) (PSuc a1) ->
   exists j, j < k /\ algo_metric j a0 a1.
 Proof.
@@ -1286,7 +1269,7 @@ Proof.
   hauto lq:on rew:off use:EJoin.suc_inj solve+:lia.
 Qed.
 
-Lemma algo_metric_bind n k p0 (A0 : PTm n) B0 p1 A1 B1  :
+Lemma algo_metric_bind k p0 (A0 : PTm ) B0 p1 A1 B1  :
   algo_metric k (PBind p0 A0 B0) (PBind p1 A1 B1) ->
   p0 = p1 /\ exists j, j < k /\ algo_metric j A0 A1 /\ algo_metric j B0 B1.
 Proof.
@@ -1306,15 +1289,15 @@ Proof.
 Qed.
 
 
-Lemma coqeq_complete' n k (a b : PTm n) :
+Lemma coqeq_complete' k (a b : PTm ) :
   algo_metric k a b ->
   (forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) /\
   (forall Γ A B, ishne a -> ishne b -> Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> a ∼ b /\ exists C, Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ∈ C /\ Γ ⊢ b ∈ C).
 Proof.
-  move : k n a b.
+  move : k a b.
   elim /Wf_nat.lt_wf_ind.
-  move => n ih.
-  move => k a b /[dup] h /algo_metric_case. move : k a b h => [:hstepL].
+  move => k ih.
+  move => a b /[dup] h /algo_metric_case. move : k a b h => [:hstepL].
   move => k a b h.
   (* Cases where a and b can take steps *)
   case; cycle 1.
diff --git a/theories/soundness.v b/theories/soundness.v
index 8bfeedf..7f86852 100644
--- a/theories/soundness.v
+++ b/theories/soundness.v
@@ -3,13 +3,13 @@ Require Import fp_red logrel typing.
 From Hammer Require Import Tactics.
 
 Theorem fundamental_theorem :
-  (forall n (Γ : fin n -> PTm n), ⊢ Γ -> ⊨ Γ) /\
-  (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> Γ ⊨ a ∈ A)  /\
-  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> Γ ⊨ a ≡ b ∈ A) /\
-  (forall n Γ (a b : PTm n), Γ ⊢ a ≲ b -> Γ ⊨ a ≲ b).
+  (forall Γ, ⊢ Γ -> ⊨ Γ) /\
+  (forall Γ a A, Γ ⊢ a ∈ A -> Γ ⊨ a ∈ A)  /\
+  (forall Γ a b A, Γ ⊢ a ≡ b ∈ A -> Γ ⊨ a ≡ b ∈ A) /\
+  (forall Γ a b, Γ ⊢ a ≲ b -> Γ ⊨ a ≲ b).
   apply wt_mutual; eauto with sem; [hauto l:on use:SE_Pair].
   Unshelve. all : exact 0.
 Qed.
 
-Lemma synsub_to_usub : forall n Γ (a b : PTm n), Γ ⊢ a ≲ b -> SN a /\ SN b /\ Sub.R a b.
+Lemma synsub_to_usub : forall Γ (a b : PTm), Γ ⊢ a ≲ b -> SN a /\ SN b /\ Sub.R a b.
 Proof. hauto lq:on rew:off use:fundamental_theorem, SemLEq_SN_Sub. Qed.
diff --git a/theories/structural.v b/theories/structural.v
index 10d6583..ccb4c6f 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -794,18 +794,17 @@ Proof.
       Unshelve. all: exact 0.
 Qed.
 
-Lemma Var_Inv Γ (i : nat) B A :
+Lemma Var_Inv Γ (i : nat) A :
   Γ ⊢ VarPTm i ∈ A ->
-  lookup i Γ B ->
-  ⊢ Γ /\ Γ ⊢ B ≲ A.
+  ⊢ Γ /\ exists B, lookup i Γ B /\ Γ ⊢ B ≲ A.
 Proof.
   move E : (VarPTm i) => u hu.
   move : i E.
   elim : Γ u A / hu=>//=.
-  - move => i Γ A hΓ hi i0 [?]. subst.
-    repeat split => //=.
-    have ? : A = B  by eauto using lookup_deter. subst.
-    have h : Γ ⊢ VarPTm i ∈ B by eauto using T_Var.
+  - move => i Γ A hΓ hi i0 [*]. subst.
+    split => //.
+    exists A. split => //.
+    have h : Γ ⊢ VarPTm i ∈ A by eauto using T_Var.
     eapply regularity in h.
     move : h => [i0]?.
     apply : Su_Eq. apply E_Refl; eassumption.

From 3a17548a7d0e7cfdc1a9e3f2d29e760ce7d0cbca Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 16:01:28 -0500
Subject: [PATCH 11/24] Minor

---
 theories/algorithmic.v | 12 ++++++------
 1 file changed, 6 insertions(+), 6 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index a52ac7e..bd0233c 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1296,18 +1296,18 @@ Lemma coqeq_complete' k (a b : PTm ) :
 Proof.
   move : k a b.
   elim /Wf_nat.lt_wf_ind.
-  move => k ih.
-  move => a b /[dup] h /algo_metric_case. move : k a b h => [:hstepL].
-  move => k a b h.
+  move => n ih.
+  move => a b /[dup] h /algo_metric_case. move : a b h => [:hstepL].
+  move => a b h.
   (* Cases where a and b can take steps *)
   case; cycle 1.
-  move : k a b h.
+  move : ih a b h.
   abstract : hstepL. qauto l:on use:HRed.preservation, CE_HRedL, hred_hne.
   move  /algo_metric_sym /algo_metric_case : (h).
   case; cycle 1.
   move {ih}. move /algo_metric_sym : h.
-  move : hstepL => /[apply].
-  hauto lq:on use:coqeq_symmetric_mutual, algo_metric_sym.
+  move : hstepL; repeat move /[apply].
+  best use:coqeq_symmetric_mutual, algo_metric_sym.
   (* Cases where a and b can't take wh steps *)
   move {hstepL}.
   move : k a b h. move => [:hnfneu].

From 5ac3b21331ff4d6b53c73e43944acb01fd12d589 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 21:09:03 -0500
Subject: [PATCH 12/24] Refactor the correctness proof of coquand's algorithm

---
 theories/algorithmic.v | 114 +++++++++++++++++++----------------------
 1 file changed, 54 insertions(+), 60 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index bd0233c..787adb1 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1305,13 +1305,15 @@ Proof.
   abstract : hstepL. qauto l:on use:HRed.preservation, CE_HRedL, hred_hne.
   move  /algo_metric_sym /algo_metric_case : (h).
   case; cycle 1.
-  move {ih}. move /algo_metric_sym : h.
-  move : hstepL; repeat move /[apply].
-  best use:coqeq_symmetric_mutual, algo_metric_sym.
+  move /algo_metric_sym : h.
+  move : hstepL ih => /[apply]/[apply].
+  repeat move /[apply].
+  move => hstepL.
+  hauto lq:on use:coqeq_symmetric_mutual, algo_metric_sym.
   (* Cases where a and b can't take wh steps *)
   move {hstepL}.
-  move : k a b h. move => [:hnfneu].
-  move => k a b h.
+  move : a b h. move => [:hnfneu].
+  move => a b h.
   case => fb; case => fa.
   - split; last by sfirstorder use:hf_not_hne.
     move {hnfneu}.
@@ -1397,11 +1399,9 @@ Proof.
           by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia.
         have {}hA0 : Γ ⊢ A0 ∈ PUniv (max i0 i1) by eauto using T_Conv.
         have {}hA1 : Γ ⊢ A1 ∈ PUniv (max i0 i1) by eauto using T_Conv.
-        have {}hB0 : funcomp (ren_PTm shift) (scons A0 Γ) ⊢
-                       B0 ∈ PUniv (max i0 i1)
+        have {}hB0 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1)
           by hauto lq:on use:T_Univ_Raise solve+:lia.
-        have {}hB1 : funcomp (ren_PTm shift) (scons A1 Γ) ⊢
-                       B1 ∈ PUniv (max i0 i1)
+        have {}hB1 : (cons A1 Γ) ⊢ B1 ∈ PUniv (max i0 i1)
           by hauto lq:on use:T_Univ_Raise solve+:lia.
 
         have ihA : A0 ⇔ A1 by hauto l:on.
@@ -1477,8 +1477,7 @@ Proof.
         move : ih h0 h2;do!move/[apply].
         move => h. apply : CE_HRed; eauto using rtc_refl.
         by constructor.
-  - move : k a b h fb fa. abstract : hnfneu.
-    move => k.
+  - move : a b h fb fa. abstract : hnfneu.
     move => + b.
     case : b => //=.
     (* NeuAbs *)
@@ -1491,14 +1490,14 @@ Proof.
         by hauto l:on use:regularity.
       have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j
           by qauto l:on use:Bind_Inv.
-      have hΓ : ⊢ funcomp (ren_PTm shift) (scons A2 Γ) by eauto using Wff_Cons'.
-      set Δ := funcomp _ _ in hΓ.
+      have hΓ : ⊢ cons A2 Γ by eauto using Wff_Cons'.
+      set Δ := cons _ _ in hΓ.
       have {}hu : Δ ⊢ PApp (ren_PTm shift u) (VarPTm var_zero) ∈ B2.
       apply : T_App'; cycle 1. apply : weakening_wt' => //=; eauto.
       reflexivity.
       rewrite -/ren_PTm.
-      by apply T_Var.
-      rewrite -/ren_PTm. by asimpl.
+      apply T_Var; eauto using here.
+      rewrite -/ren_PTm. by asimpl; rewrite subst_scons_id.
       move /Abs_Pi_Inv in ha.
       move /algo_metric_neu_abs /(_ nea) : halg.
       move => [j0][hj0]halg.
@@ -1542,7 +1541,6 @@ Proof.
         hauto lq:on rew:off use:REReds.hne_app_inv, REReds.zero_inv.
       * move => >/algo_metric_join. clear.
         hauto lq:on rew:off use:REReds.hne_proj_inv, REReds.zero_inv.
-      * sfirstorder use:T_Bot_Imp.
       * move => >/algo_metric_join. clear.
         move => h _ h2. exfalso.
         hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv.
@@ -1552,27 +1550,28 @@ Proof.
       * hauto lq:on rew:off use:REReds.var_inv, REReds.suc_inv.
       * hauto lq:on rew:off use:REReds.hne_app_inv, REReds.suc_inv.
       * hauto lq:on rew:off use:REReds.hne_proj_inv, REReds.suc_inv.
-      * sfirstorder use:T_Bot_Imp.
       * hauto q:on use:REReds.hne_ind_inv, REReds.suc_inv.
   - move {ih}.
     move /algo_metric_sym in h.
     qauto l:on use:coqeq_symmetric_mutual.
   - move {hnfneu}.
-
     (* Clear out some trivial cases *)
-    suff : (forall (Γ : fin k -> PTm k) (A B : PTm k), Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> a ∼ b /\ (exists C : PTm k, Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ∈ C /\ Γ ⊢ b ∈ C)).
+    suff : (forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> a ∼ b /\ (exists C, Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ∈ C /\ Γ ⊢ b ∈ C)).
     move => h0.
     split. move => *. apply : CE_HRed; eauto using rtc_refl. apply CE_NeuNeu. by firstorder.
     by firstorder.
 
     case : a h fb fa => //=.
     + case : b => //=; move => > /algo_metric_join.
-      * move /DJoin.var_inj => ?. subst. qauto l:on use:Var_Inv, T_Var,CE_VarCong.
+      * move /DJoin.var_inj => i _ _. subst.
+        move => Γ A B /Var_Inv [? [B0 [h0 h1]]].
+        move /Var_Inv => [_ [C0 [h2 h3]]].
+        have ? : B0 = C0 by eauto using lookup_deter. subst.
+        sfirstorder use:T_Var.
       * clear => ? ? _. exfalso.
         hauto l:on use:REReds.var_inv, REReds.hne_app_inv.
       * clear => ? ? _. exfalso.
         hauto l:on use:REReds.var_inv, REReds.hne_proj_inv.
-      * sfirstorder use:T_Bot_Imp.
       * clear => ? ? _. exfalso.
         hauto q:on use:REReds.var_inv, REReds.hne_ind_inv.
     + case : b => //=;
@@ -1629,7 +1628,6 @@ Proof.
         sfirstorder use:bind_inst.
       * clear => ? ? ?. exfalso.
         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
-      * sfirstorder use:T_Bot_Imp.
       * clear => ? ? ?. exfalso.
         hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv.
     + case : b => //=.
@@ -1691,10 +1689,8 @@ Proof.
                move /E_Symmetric in haE.
                move /regularity_sub0 in hSu21.
                sfirstorder use:bind_inst.
-      * sfirstorder use:T_Bot_Imp.
       * move => > /algo_metric_join; clear => ? ? ?. exfalso.
         hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv.
-    + sfirstorder use:T_Bot_Imp.
     (* ind ind case *)
     + move => P a0 b0 c0.
       case : b => //=;
@@ -1705,7 +1701,6 @@ Proof.
       * hauto q:on use:REReds.hne_ind_inv, REReds.var_inv.
       * hauto q:on use:REReds.hne_ind_inv, REReds.hne_app_inv.
       * hauto q:on use:REReds.hne_ind_inv, REReds.hne_proj_inv.
-      * sfirstorder use:T_Bot_Imp.
       * move => P1 a1 b1 c1 /[dup] halg /algo_metric_ind + h0 h1.
         move /(_ h1 h0).
         move => [j][hj][hP][ha][hb]hc Γ A B hL hR.
@@ -1716,7 +1711,7 @@ Proof.
         move : T_Univ_Raise wtP0; do!move/[apply]. move => wtP0.
         move : T_Univ_Raise wtP1; do!move/[apply]. move => wtP1.
         have {}ihP : P ⇔ P1 by qauto l:on.
-        set Δ := funcomp _ _ in wtP0, wtP1, wtc0, wtc1.
+        set Δ := cons _ _ in wtP0, wtP1, wtc0, wtc1.
         have wfΔ :⊢ Δ by hauto l:on use:wff_mutual.
         have hPE : Δ ⊢ P ≡ P1 ∈ PUniv (max iP0 iP1)
           by hauto l:on use:coqeq_sound_mutual.
@@ -1730,7 +1725,7 @@ Proof.
         have {}ihb : b0 ⇔ b1 by hauto l:on.
         have hPSig : Γ ⊢ PBind PSig PNat P ≡ PBind PSig PNat P1 ∈ PUniv (Nat.max iP0 iP1) by eauto with wt.
         set T := ren_PTm shift _ in wtc0.
-        have : funcomp (ren_PTm shift) (scons P Δ) ⊢ c1 ∈ T.
+        have : (cons P Δ) ⊢ c1 ∈ T.
         apply : T_Conv; eauto. apply : ctx_eq_subst_one; eauto with wt.
         apply : Su_Eq; eauto.
         subst T. apply : weakening_su; eauto.
@@ -1738,9 +1733,8 @@ Proof.
         hauto l:on use:wff_mutual.
         apply morphing_ext. set x := funcomp _ _.
         have -> : x = funcomp (ren_PTm shift) VarPTm by asimpl.
-        apply : morphing_ren; eauto using renaming_shift.
-        apply : renaming_shift; eauto. by apply morphing_id.
-        apply T_Suc. by apply T_Var. subst T => {}wtc1.
+        apply : morphing_ren; eauto using renaming_shift. by apply morphing_id.
+        apply T_Suc. apply T_Var => //=. apply here. subst T => {}wtc1.
         have {}ihc : c0 ⇔ c1 by qauto l:on.
         move => [:ih].
         split. abstract : ih. move : h0 h1 ihP iha ihb ihc. clear. sauto lq:on.
@@ -1754,11 +1748,11 @@ Proof.
         move : hEInd. clear. hauto l:on use:regularity.
 Qed.
 
-Lemma coqeq_sound : forall n Γ (a b : PTm n) A,
+Lemma coqeq_sound : forall Γ (a b : PTm) A,
     Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b -> Γ ⊢ a ≡ b ∈ A.
 Proof. sfirstorder use:coqeq_sound_mutual. Qed.
 
-Lemma coqeq_complete n Γ (a b A : PTm n) :
+Lemma coqeq_complete Γ (a b A : PTm) :
   Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
 Proof.
   move => h.
@@ -1766,7 +1760,7 @@ Proof.
   eapply fundamental_theorem in h.
   move /logrel.SemEq_SN_Join : h.
   move => h.
-  have : exists va vb : PTm n,
+  have : exists va vb : PTm,
          rtc LoRed.R a va /\
          rtc LoRed.R b vb /\ nf va /\ nf vb /\ EJoin.R va vb
       by hauto l:on use:DJoin.standardization_lo.
@@ -1780,7 +1774,7 @@ Qed.
 
 Reserved Notation "a ≪ b" (at level 70).
 Reserved Notation "a ⋖ b" (at level 70).
-Inductive CoqLEq {n} : PTm n -> PTm n -> Prop :=
+Inductive CoqLEq : PTm -> PTm -> Prop :=
 | CLE_UnivCong i j :
   i <= j ->
   (* -------------------------- *)
@@ -1805,7 +1799,7 @@ Inductive CoqLEq {n} : PTm n -> PTm n -> Prop :=
   a0 ∼ a1 ->
   a0 ⋖ a1
 
-with CoqLEq_R {n} : PTm n -> PTm n -> Prop :=
+with CoqLEq_R : PTm -> PTm -> Prop :=
 | CLE_HRed a a' b b'  :
   rtc HRed.R a a' ->
   rtc HRed.R b b' ->
@@ -1819,10 +1813,10 @@ Scheme coqleq_ind := Induction for CoqLEq Sort Prop
 
 Combined Scheme coqleq_mutual from coqleq_ind, coqleq_r_ind.
 
-Definition salgo_metric {n} k (a b : PTm n) :=
+Definition salgo_metric k (a b : PTm ) :=
   exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ nf va /\ nf vb /\ ESub.R va vb /\ size_PTm va + size_PTm vb + i + j <= k.
 
-Lemma salgo_metric_algo_metric n k (a b : PTm n) :
+Lemma salgo_metric_algo_metric k (a b : PTm) :
   ishne a \/ ishne b ->
   salgo_metric k a b ->
   algo_metric k a b.
@@ -1840,13 +1834,13 @@ Proof.
   hauto lq:on use:Sub1.hne_refl.
 Qed.
 
-Lemma coqleq_sound_mutual : forall n,
-    (forall (a b : PTm n), a ⋖ b -> forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> Γ ⊢ a ≲ b ) /\
-    (forall (a b : PTm n), a ≪ b -> forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> Γ ⊢ a ≲ b ).
+Lemma coqleq_sound_mutual :
+    (forall (a b : PTm), a ⋖ b -> forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> Γ ⊢ a ≲ b ) /\
+    (forall (a b : PTm), a ≪ b -> forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> Γ ⊢ a ≲ b ).
 Proof.
   apply coqleq_mutual.
   - hauto lq:on use:wff_mutual ctrs:LEq.
-  - move => n A0 A1 B0 B1 hA ihA hB ihB Γ i.
+  - move => A0 A1 B0 B1 hA ihA hB ihB Γ i.
     move /Bind_Univ_Inv => [hA0]hB0 /Bind_Univ_Inv [hA1]hB1.
     have hlA  : Γ ⊢ A1 ≲ A0 by sfirstorder.
     have hΓ : ⊢ Γ by sfirstorder use:wff_mutual.
@@ -1854,7 +1848,7 @@ Proof.
     by apply : Su_Pi; eauto using E_Refl, Su_Eq.
     apply : Su_Pi; eauto using E_Refl, Su_Eq.
     apply : ihB; eauto using ctx_eq_subst_one.
-  - move => n A0 A1 B0 B1 hA ihA hB ihB Γ i.
+  - move => A0 A1 B0 B1 hA ihA hB ihB Γ i.
     move /Bind_Univ_Inv => [hA0]hB0 /Bind_Univ_Inv [hA1]hB1.
     have hlA  : Γ ⊢ A0 ≲ A1 by sfirstorder.
     have hΓ : ⊢ Γ by sfirstorder use:wff_mutual.
@@ -1864,14 +1858,14 @@ Proof.
     apply : Su_Sig; eauto using E_Refl, Su_Eq.
   - sauto lq:on use:coqeq_sound_mutual, Su_Eq.
   - sauto lq:on use:coqeq_sound_mutual, Su_Eq.
-  - move => n a a' b b' ? ? ? ih Γ i ha hb.
+  - move => a a' b b' ? ? ? ih Γ i ha hb.
     have /Su_Eq ? : Γ ⊢ a ≡ a' ∈ PUniv i by sfirstorder use:HReds.ToEq.
     have /E_Symmetric /Su_Eq ? : Γ ⊢ b ≡ b' ∈ PUniv i by sfirstorder use:HReds.ToEq.
     suff : Γ ⊢ a' ≲ b' by eauto using Su_Transitive.
     eauto using HReds.preservation.
 Qed.
 
-Lemma salgo_metric_case n k (a b : PTm n) :
+Lemma salgo_metric_case k (a b : PTm ) :
   salgo_metric k a b ->
   (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ salgo_metric k' a' b /\ k' < k.
 Proof.
@@ -1899,7 +1893,7 @@ Proof.
     + sfirstorder use:ne_hne.
 Qed.
 
-Lemma CLE_HRedL n (a a' b : PTm n)  :
+Lemma CLE_HRedL (a a' b : PTm )  :
   HRed.R a a' ->
   a' ≪ b ->
   a ≪ b.
@@ -1907,7 +1901,7 @@ Proof.
   hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
 Qed.
 
-Lemma CLE_HRedR n (a a' b : PTm n)  :
+Lemma CLE_HRedR (a a' b : PTm)  :
   HRed.R a a' ->
   b ≪ a' ->
   b ≪ a.
@@ -1916,7 +1910,7 @@ Proof.
 Qed.
 
 
-Lemma algo_metric_caseR n k (a b : PTm n) :
+Lemma algo_metric_caseR k (a b : PTm) :
   salgo_metric k a b ->
   (ishf b \/ ishne b) \/ exists k' b', HRed.R b b' /\ salgo_metric k' a b' /\ k' < k.
 Proof.
@@ -1944,7 +1938,7 @@ Proof.
     + sfirstorder use:ne_hne.
 Qed.
 
-Lemma salgo_metric_sub n k (a b : PTm n) :
+Lemma salgo_metric_sub k (a b : PTm) :
   salgo_metric k a b ->
   Sub.R a b.
 Proof.
@@ -1958,7 +1952,7 @@ Proof.
   rewrite /Sub.R. exists va', vb'. sfirstorder use:@relations.rtc_transitive.
 Qed.
 
-Lemma salgo_metric_pi n k (A0 : PTm n) B0 A1 B1  :
+Lemma salgo_metric_pi k (A0 : PTm) B0 A1 B1  :
   salgo_metric k (PBind PPi A0 B0) (PBind PPi A1 B1) ->
   exists j, j < k /\ salgo_metric j A1 A0 /\ salgo_metric j B0 B1.
 Proof.
@@ -1971,7 +1965,7 @@ Proof.
   hauto qb:on solve+:lia.
 Qed.
 
-Lemma salgo_metric_sig n k (A0 : PTm n) B0 A1 B1  :
+Lemma salgo_metric_sig k (A0 : PTm) B0 A1 B1  :
   salgo_metric k (PBind PSig A0 B0) (PBind PSig A1 B1) ->
   exists j, j < k /\ salgo_metric j A0 A1 /\ salgo_metric j B0 B1.
 Proof.
@@ -1984,16 +1978,16 @@ Proof.
   hauto qb:on solve+:lia.
 Qed.
 
-Lemma coqleq_complete' n k (a b : PTm n) :
+Lemma coqleq_complete' k (a b : PTm ) :
   salgo_metric k a b -> (forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> a ≪ b).
 Proof.
-  move : k n a b.
+  move : k a b.
   elim /Wf_nat.lt_wf_ind.
   move => n ih.
-  move => k a b /[dup] h /salgo_metric_case.
+  move => a b /[dup] h /salgo_metric_case.
   (* Cases where a and b can take steps *)
   case; cycle 1.
-  move : k a b h.
+  move : a b h.
   qauto l:on use:HRed.preservation, CLE_HRedL, hred_hne.
   case /algo_metric_caseR : (h); cycle 1.
   qauto l:on use:HRed.preservation, CLE_HRedR, hred_hne.
@@ -2013,7 +2007,7 @@ Proof.
            have ihA : A0 ≪ A1 by hauto l:on.
            econstructor; eauto using E_Refl; constructor=> //.
            have ihA' : Γ ⊢ A0 ≲ A1 by hauto l:on use:coqleq_sound_mutual.
-           suff : funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B1 ∈ PUniv i
+           suff : (cons A0 Γ) ⊢ B1 ∈ PUniv i
              by hauto l:on.
            eauto using ctx_eq_subst_one.
         ** move /salgo_metric_sig.
@@ -2022,7 +2016,7 @@ Proof.
            have ihA : A1 ≪ A0 by hauto l:on.
            econstructor; eauto using E_Refl; constructor=> //.
            have ihA' : Γ ⊢ A1 ≲ A0 by hauto l:on use:coqleq_sound_mutual.
-           suff : funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ∈ PUniv i
+           suff : (cons A1 Γ) ⊢ B0 ∈ PUniv i
              by hauto l:on.
            eauto using ctx_eq_subst_one.
       * hauto lq:on use:salgo_metric_sub, Sub.bind_univ_noconf.
@@ -2083,7 +2077,7 @@ Proof.
     eapply coqeq_complete'; eauto.
 Qed.
 
-Lemma coqleq_complete n Γ (a b : PTm n) :
+Lemma coqleq_complete Γ (a b : PTm) :
   Γ ⊢ a ≲ b -> a ≪ b.
 Proof.
   move => h.
@@ -2091,7 +2085,7 @@ Proof.
   eapply fundamental_theorem in h.
   move /logrel.SemLEq_SN_Sub : h.
   move => h.
-  have : exists va vb : PTm n,
+  have : exists va vb : PTm ,
          rtc LoRed.R a va /\
          rtc LoRed.R b vb /\ nf va /\ nf vb /\ ESub.R va vb
       by hauto l:on use:Sub.standardization_lo.
@@ -2102,10 +2096,10 @@ Proof.
   hauto lq:on solve+:lia.
 Qed.
 
-Lemma coqleq_sound : forall n Γ (a b : PTm n) i j,
+Lemma coqleq_sound : forall Γ (a b : PTm) i j,
     Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv j -> a ≪ b -> Γ ⊢ a ≲ b.
 Proof.
-  move => n Γ a b i j.
+  move => Γ a b i j.
   have [*] : i <= i + j /\ j <= i + j by lia.
   have : Γ ⊢ a ∈ PUniv (i + j) /\ Γ ⊢ b ∈ PUniv (i + j)
     by sfirstorder use:T_Univ_Raise.

From 36d1f95d6524b7f7934172af76e8231b9b54496f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 21:16:42 -0500
Subject: [PATCH 13/24] Move HRed to the common .v file

---
 theories/algorithmic.v | 25 -------------------------
 theories/common.v      | 27 +++++++++++++++++++++++++++
 2 files changed, 27 insertions(+), 25 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 787adb1..8a9146a 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -8,31 +8,6 @@ Require Import Coq.Logic.FunctionalExtensionality.
 Require Import Cdcl.Itauto.
 
 Module HRed.
-  Inductive R : PTm -> PTm ->  Prop :=
-  (****************** Beta ***********************)
-  | AppAbs a b :
-    R (PApp (PAbs a) b)  (subst_PTm (scons b VarPTm) a)
-
-  | ProjPair p a b :
-    R (PProj p (PPair a b)) (if p is PL then a else b)
-
-  | IndZero P b c :
-    R (PInd P PZero b c) b
-
-  | IndSuc P a b c :
-    R (PInd P (PSuc a) b c) (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
-
-  (*************** Congruence ********************)
-  | AppCong a0 a1 b :
-    R a0 a1 ->
-    R (PApp a0 b) (PApp a1 b)
-  | ProjCong p a0 a1 :
-    R a0 a1 ->
-    R (PProj p a0) (PProj p a1)
-  | IndCong P a0 a1 b c :
-    R a0 a1 ->
-    R (PInd P a0 b c) (PInd P a1 b c).
-
   Lemma ToRRed (a b : PTm) : HRed.R a b -> RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
 
diff --git a/theories/common.v b/theories/common.v
index 5095fef..9c6b508 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -143,3 +143,30 @@ Proof.
   rewrite -{2}E.
   apply ext_PTm. case => //=.
 Qed.
+
+Module HRed.
+    Inductive R : PTm -> PTm ->  Prop :=
+  (****************** Beta ***********************)
+  | AppAbs a b :
+    R (PApp (PAbs a) b)  (subst_PTm (scons b VarPTm) a)
+
+  | ProjPair p a b :
+    R (PProj p (PPair a b)) (if p is PL then a else b)
+
+  | IndZero P b c :
+    R (PInd P PZero b c) b
+
+  | IndSuc P a b c :
+    R (PInd P (PSuc a) b c) (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
+
+  (*************** Congruence ********************)
+  | AppCong a0 a1 b :
+    R a0 a1 ->
+    R (PApp a0 b) (PApp a1 b)
+  | ProjCong p a0 a1 :
+    R a0 a1 ->
+    R (PProj p a0) (PProj p a1)
+  | IndCong P a0 a1 b c :
+    R a0 a1 ->
+    R (PInd P a0 b c) (PInd P a1 b c).
+End HRed.

From 87f6dcd870883f29892b1d9f81c9d57f1262dcc4 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Mar 2025 23:46:41 -0500
Subject: [PATCH 14/24] Prove the soundness of the computable equality

---
 theories/common.v             |   5 +
 theories/executable.v         | 340 +++++++++++++++++++++++++++-------
 theories/executable_correct.v | 162 ++++++++++++++++
 3 files changed, 440 insertions(+), 67 deletions(-)
 create mode 100644 theories/executable_correct.v

diff --git a/theories/common.v b/theories/common.v
index 9c6b508..3ea15d6 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -1,4 +1,7 @@
 Require Import Autosubst2.unscoped Autosubst2.syntax Autosubst2.core ssreflect.
+From Equations Require Import Equations.
+Derive NoConfusion for nat PTag BTag PTm.
+Derive EqDec for BTag PTag PTm.
 From Ltac2 Require Ltac2.
 Import Ltac2.Notations.
 Import Ltac2.Control.
@@ -169,4 +172,6 @@ Module HRed.
   | IndCong P a0 a1 b c :
     R a0 a1 ->
     R (PInd P a0 b c) (PInd P a1 b c).
+
+    Definition nf a := forall b, ~ R a b.
 End HRed.
diff --git a/theories/executable.v b/theories/executable.v
index f773fa6..5de3355 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -1,58 +1,145 @@
 From Equations Require Import Equations.
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
-  common typing preservation admissible fp_red structural soundness.
-Require Import algorithmic.
-From stdpp Require Import relations (rtc(..), nsteps(..)).
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 Require Import ssreflect ssrbool.
+Import Logic (inspect).
 
-Inductive algo_dom {n} : PTm n -> PTm n -> Prop :=
+Require Import ssreflect ssrbool.
+From Hammer Require Import Tactics.
+
+Definition tm_nonconf (a b : PTm) : bool :=
+  match a, b with
+  | PAbs _, _ => ishne b || isabs b
+  | _, PAbs _ => ishne a
+  | VarPTm _, VarPTm _ => true
+  | PPair _ _, _ => ishne b || ispair b
+  | _, PPair _ _ => ishne a
+  | PZero, PZero => true
+  | PSuc _, PSuc _ => true
+  | PApp _ _, PApp _ _ => ishne a && ishne b
+  | PProj _ _, PProj _ _ => ishne a && ishne b
+  | PInd _ _ _ _, PInd _ _ _ _ => ishne a && ishne b
+  | PNat, PNat => true
+  | PUniv _, PUniv _ => true
+  | PBind _ _ _, PBind _ _ _ => true
+  | _,_=> false
+  end.
+
+Definition tm_conf (a b : PTm) := ~~ tm_nonconf a b.
+
+Inductive eq_view : PTm -> PTm -> Type :=
+| V_AbsAbs a b :
+  eq_view (PAbs a) (PAbs b)
+| V_AbsNeu a b :
+  ~~ isabs b ->
+  eq_view (PAbs a) b
+| V_NeuAbs a b :
+  ~~ isabs a ->
+  eq_view a (PAbs b)
+| V_VarVar i j :
+  eq_view (VarPTm i) (VarPTm j)
+| V_PairPair a0 b0 a1 b1 :
+  eq_view (PPair a0 b0) (PPair a1 b1)
+| V_PairNeu a0 b0 u :
+  ~~ ispair u ->
+  eq_view (PPair a0 b0) u
+| V_NeuPair u a1 b1 :
+  ~~ ispair u ->
+  eq_view u (PPair a1 b1)
+| V_ZeroZero :
+  eq_view PZero PZero
+| V_SucSuc a b :
+  eq_view (PSuc a) (PSuc b)
+| V_AppApp u0 b0 u1 b1 :
+  eq_view (PApp u0 b0) (PApp u1 b1)
+| V_ProjProj p0 u0 p1 u1 :
+  eq_view (PProj p0 u0) (PProj p1 u1)
+| V_IndInd P0 u0 b0 c0  P1 u1 b1 c1 :
+  eq_view (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
+| V_NatNat :
+  eq_view PNat PNat
+| V_BindBind p0 A0 B0 p1 A1 B1 :
+  eq_view (PBind p0 A0 B0) (PBind p1 A1 B1)
+| V_UnivUniv i j :
+  eq_view (PUniv i) (PUniv j)
+| V_Others a b :
+  eq_view a b.
+
+Equations tm_to_eq_view (a b : PTm) : eq_view a b :=
+  tm_to_eq_view (PAbs a) (PAbs b) := V_AbsAbs a b;
+  tm_to_eq_view (PAbs a) b := V_AbsNeu a b _;
+  tm_to_eq_view a (PAbs b) := V_NeuAbs a b _;
+  tm_to_eq_view (VarPTm i) (VarPTm j) := V_VarVar i j;
+  tm_to_eq_view (PPair a0 b0) (PPair a1 b1) := V_PairPair a0 b0 a1 b1;
+  tm_to_eq_view (PPair a0 b0) u := V_PairNeu a0 b0 u _;
+  tm_to_eq_view u (PPair a1 b1) := V_NeuPair u a1 b1 _;
+  tm_to_eq_view PZero PZero := V_ZeroZero;
+  tm_to_eq_view (PSuc a) (PSuc b) := V_SucSuc a b;
+  tm_to_eq_view (PApp a0 b0) (PApp a1 b1) := V_AppApp a0 b0 a1 b1;
+  tm_to_eq_view (PProj p0 b0) (PProj p1 b1) := V_ProjProj p0 b0 p1 b1;
+  tm_to_eq_view (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1) := V_IndInd P0 u0 b0 c0 P1 u1 b1 c1;
+  tm_to_eq_view PNat PNat := V_NatNat;
+  tm_to_eq_view (PUniv i) (PUniv j) := V_UnivUniv i j;
+  tm_to_eq_view (PBind p0 A0 B0)  (PBind p1 A1 B1) := V_BindBind p0 A0 B0 p1 A1 B1;
+  tm_to_eq_view a b := V_Others a b.
+
+Inductive algo_dom : PTm -> PTm -> Prop :=
 | A_AbsAbs a b :
-  algo_dom a b ->
+  algo_dom_r a b ->
   (* --------------------- *)
   algo_dom (PAbs a) (PAbs b)
 
 | A_AbsNeu a u :
   ishne u ->
-  algo_dom a (PApp (ren_PTm shift u) (VarPTm var_zero)) ->
+  algo_dom_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) ->
   (* --------------------- *)
   algo_dom (PAbs a) u
 
 | A_NeuAbs a u :
   ishne u ->
-  algo_dom (PApp (ren_PTm shift u) (VarPTm var_zero)) a ->
+  algo_dom_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a ->
   (* --------------------- *)
   algo_dom u (PAbs a)
 
 | A_PairPair a0 a1 b0 b1 :
-  algo_dom a0 a1 ->
-  algo_dom b0 b1 ->
+  algo_dom_r a0 a1 ->
+  algo_dom_r b0 b1 ->
   (* ---------------------------- *)
   algo_dom (PPair a0 b0) (PPair a1 b1)
 
 | A_PairNeu a0 a1 u :
   ishne u ->
-  algo_dom a0 (PProj PL u) ->
-  algo_dom a1 (PProj PR u) ->
+  algo_dom_r a0 (PProj PL u) ->
+  algo_dom_r a1 (PProj PR u) ->
   (* ----------------------- *)
   algo_dom (PPair a0 a1) u
 
 | A_NeuPair a0 a1 u :
   ishne u ->
-  algo_dom (PProj PL u) a0 ->
-  algo_dom (PProj PR u) a1 ->
+  algo_dom_r (PProj PL u) a0 ->
+  algo_dom_r (PProj PR u) a1 ->
   (* ----------------------- *)
   algo_dom u (PPair a0 a1)
 
+| A_ZeroZero :
+  algo_dom PZero PZero
+
+| A_SucSuc a0 a1 :
+  algo_dom_r a0 a1 ->
+  algo_dom (PSuc a0) (PSuc a1)
+
 | A_UnivCong i j :
   (* -------------------------- *)
   algo_dom (PUniv i) (PUniv j)
 
 | A_BindCong p0 p1 A0 A1 B0 B1 :
-  algo_dom A0 A1 ->
-  algo_dom B0 B1 ->
+  algo_dom_r A0 A1 ->
+  algo_dom_r B0 B1 ->
   (* ---------------------------- *)
   algo_dom (PBind p0 A0 B0) (PBind p1 A1 B1)
 
+| A_NatCong :
+  algo_dom PNat PNat
+
 | A_VarCong i j :
   (* -------------------------- *)
   algo_dom (VarPTm i) (VarPTm j)
@@ -68,78 +155,197 @@ Inductive algo_dom {n} : PTm n -> PTm n -> Prop :=
   ishne u0 ->
   ishne u1 ->
   algo_dom u0 u1 ->
-  algo_dom a0 a1 ->
+  algo_dom_r a0 a1 ->
   (* ------------------------- *)
   algo_dom (PApp u0 a0) (PApp u1 a1)
 
+| A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 :
+  ishne u0 ->
+  ishne u1 ->
+  algo_dom_r P0 P1 ->
+  algo_dom u0 u1 ->
+  algo_dom_r b0 b1 ->
+  algo_dom_r c0 c1 ->
+  algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
+
+with algo_dom_r : PTm  -> PTm -> Prop :=
+| A_NfNf a b :
+  algo_dom a b ->
+  algo_dom_r a b
+
 | A_HRedL a a' b  :
   HRed.R a a' ->
-  algo_dom a' b ->
+  algo_dom_r a' b ->
   (* ----------------------- *)
-  algo_dom a b
+  algo_dom_r a b
 
 | A_HRedR a b b'  :
   ishne a \/ ishf a ->
   HRed.R b b' ->
-  algo_dom a b' ->
+  algo_dom_r a b' ->
   (* ----------------------- *)
-  algo_dom a b.
+  algo_dom_r a b.
 
-
-Definition fin_eq {n} (i j : fin n) : bool.
+Lemma algo_dom_hf_hne (a b : PTm) :
+  algo_dom a b ->
+  (ishf a \/ ishne a) /\ (ishf b \/ ishne b).
 Proof.
-  induction n.
-  - by exfalso.
-  - refine (match i , j with
-            | None, None => true
-            | Some i, Some j => IHn i j
-            | _, _ => false
-            end).
+  induction 1 =>//=; hauto lq:on.
+Qed.
+
+Lemma hf_no_hred (a b : PTm) :
+  ishf a ->
+  HRed.R a b ->
+  False.
+Proof. hauto l:on inv:HRed.R. Qed.
+
+Lemma hne_no_hred (a b : PTm) :
+  ishne a ->
+  HRed.R a b ->
+  False.
+Proof. elim : a b => //=; hauto l:on inv:HRed.R. Qed.
+
+Derive Signature for algo_dom algo_dom_r.
+
+Fixpoint hred (a : PTm) : option (PTm) :=
+    match a with
+    | VarPTm i => None
+    | PAbs a => None
+    | PApp (PAbs a) b => Some (subst_PTm (scons b VarPTm) a)
+    | PApp a b =>
+        match hred a with
+        | Some a => Some (PApp a b)
+        | None => None
+        end
+    | PPair a b => None
+    | PProj p (PPair a b) => if p is PL then Some a else Some b
+    | PProj p a =>
+        match hred a with
+        | Some a => Some (PProj p a)
+        | None => None
+        end
+    | PUniv i => None
+    | PBind p A B => None
+    | PNat => None
+    | PZero => None
+    | PSuc a => None
+    | PInd P PZero b c => Some b
+    | PInd P (PSuc a) b c =>
+        Some (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
+    | PInd P a b c =>
+        match hred a with
+        | Some a => Some (PInd P a b c)
+        | None => None
+        end
+    end.
+
+Lemma hred_complete (a b : PTm) :
+  HRed.R a b -> hred a = Some b.
+Proof.
+  induction 1; hauto lq:on rew:off inv:HRed.R b:on.
+Qed.
+
+Lemma hred_sound (a b : PTm):
+  hred a = Some b -> HRed.R a b.
+Proof.
+  elim : a b; hauto q:on dep:on ctrs:HRed.R.
+Qed.
+
+Lemma hred_deter (a b0 b1 : PTm) :
+  HRed.R a b0 -> HRed.R a b1 -> b0 = b1.
+Proof.
+  move /hred_complete => + /hred_complete. congruence.
+Qed.
+
+Definition fancy_hred (a : PTm) : HRed.nf a + {b | HRed.R a b}.
+  destruct (hred a) eqn:eq.
+  right. exists p. by apply hred_sound in eq.
+  left. move => b /hred_complete. congruence.
 Defined.
 
-Lemma fin_eq_dec {n} (i j : fin n) :
-  Bool.reflect (i = j) (fin_eq i j).
-Proof.
-  revert i j. induction n.
-  - destruct i.
-  - destruct i; destruct j.
-    + specialize (IHn f f0).
-      inversion IHn; subst.
-      simpl. rewrite -H.
-      apply ReflectT.
-      reflexivity.
-      simpl. rewrite -H.
-      apply ReflectF.
-      injection. tauto.
-    + by apply ReflectF.
-    + by apply ReflectF.
-    + by apply ReflectT.
-Defined.
+Ltac check_equal_triv :=
+  intros;subst;
+  lazymatch goal with
+  (* | [h : algo_dom (VarPTm _) (PAbs _) |- _] => idtac *)
+  | [h : algo_dom _ _ |- _] => try (inversion h; by subst)
+  | _ => idtac
+  end.
 
+Ltac solve_check_equal :=
+  try solve [intros *;
+  match goal with
+  | [|- _ = _] => sauto
+  | _ => idtac
+  end].
 
-Equations check_equal {n} (a b : PTm n) (h : algo_dom a b) :
+Equations check_equal (a b : PTm) (h : algo_dom a b) :
   bool by struct h :=
-  check_equal a b h with (@idP (ishne a || ishf a)) := {
-    | Bool.ReflectT _ _ => _
-    | Bool.ReflectF _ _ => _
-    }.
+  check_equal a b h with tm_to_eq_view a b :=
+  check_equal _ _ h (V_VarVar i j) := nat_eqdec i j;
+  check_equal _ _ h (V_AbsAbs a b) := check_equal_r a b ltac:(check_equal_triv);
+  check_equal _ _ h (V_AbsNeu a b h') := check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) ltac:(check_equal_triv);
+  check_equal _ _ h (V_NeuAbs a b _) := check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b ltac:(check_equal_triv);
+  check_equal _ _ h (V_PairPair a0 b0 a1 b1) :=
+    check_equal_r a0 a1 ltac:(check_equal_triv) && check_equal_r b0 b1 ltac:(check_equal_triv);
+  check_equal _ _ h (V_PairNeu a0 b0 u _) :=
+    check_equal_r a0 (PProj PL u) ltac:(check_equal_triv) && check_equal_r b0 (PProj PR u) ltac:(check_equal_triv);
+  check_equal _ _ h (V_NeuPair u a1 b1 _) :=
+    check_equal_r (PProj PL u) a1 ltac:(check_equal_triv) && check_equal_r (PProj PR u) b1 ltac:(check_equal_triv);
+  check_equal _ _ h V_NatNat := true;
+  check_equal _ _ h V_ZeroZero := true;
+  check_equal _ _ h (V_SucSuc a b) := check_equal_r a b ltac:(check_equal_triv);
+  check_equal _ _ h (V_ProjProj p0 a p1 b) :=
+    PTag_eqdec p0 p1 && check_equal a b ltac:(check_equal_triv);
+  check_equal _ _ h (V_AppApp a0 b0 a1 b1) :=
+    check_equal a0 a1 ltac:(check_equal_triv) && check_equal_r b0 b1 ltac:(check_equal_triv);
+  check_equal _ _ h (V_IndInd P0 u0 b0 c0 P1 u1 b1 c1) :=
+    check_equal_r P0 P1 ltac:(check_equal_triv) &&
+      check_equal u0 u1 ltac:(check_equal_triv) &&
+      check_equal_r b0 b1 ltac:(check_equal_triv) &&
+      check_equal_r c0 c1 ltac:(check_equal_triv);
+  check_equal _ _ h (V_UnivUniv i j) := nat_eqdec i j;
+  check_equal _ _ h (V_BindBind p0 A0 B0 p1 A1 B1) :=
+    BTag_eqdec p0 p1 && check_equal_r A0 A1 ltac:(check_equal_triv) && check_equal_r B0 B1 ltac:(check_equal_triv);
+  check_equal _ _ _ _ := false
 
+  (* check_equal a b h := false; *)
+  with check_equal_r (a b : PTm) (h0 : algo_dom_r a b) :
+  bool by struct h0 :=
+    check_equal_r a b h with (fancy_hred a) :=
+      check_equal_r a b h (inr a') := check_equal_r (proj1_sig a') b _;
+      check_equal_r a b h (inl h')  with (fancy_hred b) :=
+        | inr b' := check_equal_r a (proj1_sig b') _;
+        | inl h'' := check_equal a b _.
 
-  (* check_equal (VarPTm i) (VarPTm j) h := fin_eq i j; *)
-  (* check_equal (PAbs a) (PAbs b) h := check_equal a b _; *)
-  (* check_equal (PPair a0 b0) (PPair a1 b1) h := *)
-  (*   check_equal a0 b0 _ && check_equal a1 b1 _; *)
-  (* check_equal (PUniv i) (PUniv j) _ := _; *)
 Next Obligation.
-  simpl.
-  intros ih.
-Admitted.
+  intros.
+  inversion h; subst => //=.
+  exfalso. hauto l:on use:hred_complete unfold:HRed.nf.
+  exfalso. hauto l:on use:hred_complete unfold:HRed.nf.
+Defined.
 
-Search (Bool.reflect (is_true _) _).
-Check idP.
+Next Obligation.
+  intros.
+  destruct h.
+  exfalso. apply algo_dom_hf_hne in H0.
+  destruct H0 as [h0 h1].
+  sfirstorder use:hf_no_hred, hne_no_hred.
+  exfalso. sfirstorder.
+  assert (  b' = b'0)by eauto using hred_deter. subst.
+  apply h.
+Defined.
 
-Definition metric {n} k (a b : PTm n) :=
-  exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\
-  nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j <= k.
+Next Obligation.
+  simpl. intros.
+  inversion h; subst =>//=.
+  move {h}. hauto lq:on use:algo_dom_hf_hne, hf_no_hred, hne_no_hred, hred_sound.
+  assert (a' = a'0) by eauto using hred_deter, hred_sound. by subst.
+  exfalso. sfirstorder use:hne_no_hred, hf_no_hred.
+Defined.
 
-Search (nat -> nat -> bool).
+
+
+
+Next Obligation.
+  qauto inv:algo_dom, algo_dom_r.
+Defined.
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
new file mode 100644
index 0000000..603c91c
--- /dev/null
+++ b/theories/executable_correct.v
@@ -0,0 +1,162 @@
+From Equations Require Import Equations.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common executable algorithmic.
+Require Import ssreflect ssrbool.
+From stdpp Require Import relations (rtc(..)).
+From Hammer Require Import Tactics.
+
+Scheme algo_ind := Induction for algo_dom Sort Prop
+  with algor_ind := Induction for algo_dom_r Sort Prop.
+
+Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
+
+Lemma check_equal_pair_neu a0 a1 u neu h h'
+  : check_equal (PPair a0 a1) u (A_PairNeu a0 a1 u neu h h') = check_equal_r a0 (PProj PL u) h && check_equal_r a1 (PProj PR u) h'.
+Proof.
+  case : u neu h h' => //=; simp check_equal tm_to_eq_view.
+Qed.
+
+Lemma check_equal_neu_pair a0 a1 u neu h h'
+  : check_equal u (PPair a0 a1) (A_NeuPair a0 a1 u neu h h') = check_equal_r  (PProj PL u) a0 h && check_equal_r (PProj PR u) a1 h'.
+Proof.
+  case : u neu h h' => //=; simp check_equal tm_to_eq_view.
+Qed.
+
+Lemma check_equal_bind_bind p0 A0 B0 p1 A1 B1 h0 h1 :
+  check_equal (PBind p0 A0 B0) (PBind p1 A1 B1) (A_BindCong p0 p1 A0 A1 B0 B1 h0 h1) =
+    BTag_eqdec p0 p1 &&  check_equal_r A0 A1 h0 && check_equal_r  B0 B1 h1.
+Proof. hauto lq:on. Qed.
+
+Lemma check_equal_proj_proj p0 u0 p1 u1 neu0 neu1 h :
+  check_equal (PProj p0 u0) (PProj p1 u1) (A_ProjCong p0 p1 u0 u1 neu0 neu1 h) =
+    PTag_eqdec p0 p1 && check_equal u0 u1 h.
+Proof. hauto lq:on. Qed.
+
+Lemma check_equal_app_app u0 a0 u1 a1 hu0 hu1 hdom hdom' :
+  check_equal (PApp u0 a0) (PApp u1 a1) (A_AppCong u0 u1 a0 a1 hu0 hu1 hdom hdom') =
+    check_equal u0 u1 hdom && check_equal_r a0 a1 hdom'.
+Proof. hauto lq:on. Qed.
+
+Lemma check_equal_ind_ind P0 u0 b0 c0 P1 u1 b1 c1 neu0 neu1 domP domu domb domc :
+  check_equal (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
+    (A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 domP domu domb domc) =
+    check_equal_r P0 P1 domP && check_equal u0 u1 domu && check_equal_r b0 b1 domb && check_equal_r c0 c1 domc.
+Proof. hauto lq:on. Qed.
+
+Lemma hred_none a : HRed.nf a -> hred a = None.
+Proof.
+  destruct (hred a) eqn:eq;
+  sfirstorder use:hred_complete, hred_sound.
+Qed.
+
+Lemma check_equal_nfnf a b dom : check_equal_r a b (A_NfNf a b dom) = check_equal a b dom.
+Proof.
+  have [h0 h1] : (ishf a \/ ishne a) /\ (ishf b \/ ishne b) by hauto l:on use:algo_dom_hf_hne.
+  have [h3 h4] : hred a = None /\ hred b = None by sfirstorder use:hf_no_hred, hne_no_hred, hred_none.
+  simp check_equal.
+  destruct (fancy_hred a).
+  simp check_equal.
+  destruct (fancy_hred b).
+  simp check_equal. hauto lq:on.
+  exfalso. hauto l:on use:hred_complete.
+  exfalso. hauto l:on use:hred_complete.
+Qed.
+
+Lemma check_equal_hredl a b a' ha doma :
+  check_equal_r a b (A_HRedL a a' b ha doma) = check_equal_r a' b doma.
+Proof.
+  simp check_equal.
+  destruct (fancy_hred a).
+  - hauto q:on unfold:HRed.nf.
+  - simp check_equal.
+    destruct s as [x ?]. have ? : x = a' by eauto using hred_deter. subst.
+    simpl.
+    simp check_equal.
+    f_equal.
+    apply PropExtensionality.proof_irrelevance.
+Qed.
+
+Lemma check_equal_hredr a b b' hu r a0 :
+  check_equal_r a b (A_HRedR a b b' hu r a0) =
+    check_equal_r a b' a0.
+Proof.
+  simp check_equal.
+  destruct (fancy_hred a).
+  - rewrite check_equal_r_clause_1_equation_1.
+    destruct (fancy_hred b) as [|[b'' hb']].
+    + hauto lq:on unfold:HRed.nf.
+    + have ? : (b'' = b') by eauto using hred_deter. subst.
+      rewrite check_equal_r_clause_1_equation_1.  simpl.
+      simp check_equal. destruct (fancy_hred a). simp check_equal.
+      f_equal; apply PropExtensionality.proof_irrelevance.
+      simp check_equal. exfalso. sfirstorder use:hne_no_hred, hf_no_hred.
+  - simp check_equal. exfalso.
+    sfirstorder use:hne_no_hred, hf_no_hred.
+Qed.
+
+Lemma coqeq_neuneu u0 u1 :
+  ishne u0 -> ishne u1 -> u0 ↔ u1 -> u0 ∼ u1.
+Proof.
+  inversion 3; subst => //=.
+Qed.
+
+Lemma check_equal_sound :
+  (forall a b (h : algo_dom a b), check_equal a b h -> a ↔ b) /\
+  (forall a b (h : algo_dom_r a b), check_equal_r a b h -> a ⇔ b).
+Proof.
+  apply algo_dom_mutual.
+  - move => a b h.
+    move => h0 h1.
+    have {}h1 : check_equal_r a b h by hauto l:on rew:db:check_equal.
+    constructor. tauto.
+  - move => a u i h0 ih h.
+    apply CE_AbsNeu => //.
+    apply : ih. simp check_equal tm_to_eq_view in h.
+    have h1 : check_equal (PAbs a) u (A_AbsNeu a u i h0) = check_equal_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) h0 by  clear; case : u i h0 => //=.
+    hauto lq:on.
+  - move => a u i h ih h0.
+    apply CE_NeuAbs=>//.
+    apply ih.
+    case : u i h ih h0 => //=.
+  - move => a0 a1 b0 b1 a ha h.
+    move => h0 h1. constructor. apply ha. hauto lb:on rew:db:check_equal.
+    apply h0. hauto lb:on rew:db:check_equal.
+  - move => a0 a1 u neu h ih h' ih' he.
+    rewrite check_equal_pair_neu in he.
+    apply CE_PairNeu => //; hauto lqb:on.
+  - move => a0 a1 u i a ha a2 hb.
+    rewrite check_equal_neu_pair => *.
+    apply CE_NeuPair => //; hauto lqb:on.
+  - sfirstorder.
+  - hauto l:on use:CE_SucSuc.
+  - move => i j /sumboolP.
+    hauto lq:on use:CE_UnivCong.
+  - move => p0 p1 A0 A1 B0 B1 h0 ih0 h1 ih1 h2.
+    rewrite check_equal_bind_bind in h2.
+    move : h2.
+    move /andP => [/andP [h20 h21] h3].
+    move /sumboolP : h20 => ?. subst.
+    hauto l:on use:CE_BindCong.
+  - sfirstorder.
+  - move => i j /sumboolP ?.  subst.
+    apply : CE_NeuNeu. apply CE_VarCong.
+  - move => p0 p1 u0 u1 neu0 neu1 h ih he.
+    apply CE_NeuNeu.
+    rewrite check_equal_proj_proj in he.
+    move /andP : he => [/sumboolP ? h1]. subst.
+    sauto lq:on use:coqeq_neuneu.
+  - move => u0 u1 a0 a1 hu0 hu1 hdom ih hdom' ih' hE.
+    rewrite check_equal_app_app in hE.
+    move /andP : hE => [h0 h1].
+    sauto lq:on use:coqeq_neuneu.
+  - move => P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 domP ihP domu ihu domb ihb domc ihc.
+    rewrite check_equal_ind_ind.
+    move /andP => [/andP [/andP [h0 h1] h2 ] h3].
+    sauto lq:on use:coqeq_neuneu.
+  - move => a b dom h ih. apply : CE_HRed; eauto using rtc_refl.
+    rewrite check_equal_nfnf in ih.
+    tauto.
+  - move => a a' b ha doma ih hE.
+    rewrite check_equal_hredl in hE. sauto lq:on.
+  - move => a b b' hu r a0 ha hb. rewrite check_equal_hredr in hb.
+    sauto lq:on rew:off.
+Qed.

From 0060d3fb8612e4c0aee439812fbbbb5f790057c9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 4 Mar 2025 00:27:42 -0500
Subject: [PATCH 15/24] Factor out the rewriting lemmas

---
 theories/executable_correct.v | 25 ++++++++++++++++++++-----
 1 file changed, 20 insertions(+), 5 deletions(-)

diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 603c91c..d520857 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -9,6 +9,20 @@ Scheme algo_ind := Induction for algo_dom Sort Prop
 
 Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
 
+Lemma check_equal_abs_abs a b h : check_equal (PAbs a) (PAbs b) (A_AbsAbs a b h) = check_equal_r a b h.
+Proof. hauto l:on rew:db:check_equal. Qed.
+
+Lemma check_equal_abs_neu a u neu h : check_equal (PAbs a) u (A_AbsNeu a u neu h) = check_equal_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) h.
+Proof. case : u neu h => //=.  Qed.
+
+Lemma check_equal_neu_abs a u neu h : check_equal u (PAbs a) (A_NeuAbs a u neu h) = check_equal_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a h.
+Proof. case : u neu h => //=.  Qed.
+
+Lemma check_equal_pair_pair a0 b0 a1 b1 a h :
+  check_equal (PPair a0 b0) (PPair a1 b1) (A_PairPair a0 a1 b0 b1 a h) =
+    check_equal_r a0 a1 a && check_equal_r b0 b1 h.
+Proof. hauto l:on rew:db:check_equal. Qed.
+
 Lemma check_equal_pair_neu a0 a1 u neu h h'
   : check_equal (PPair a0 a1) u (A_PairNeu a0 a1 u neu h h') = check_equal_r a0 (PProj PL u) h && check_equal_r a1 (PProj PR u) h'.
 Proof.
@@ -105,8 +119,8 @@ Lemma check_equal_sound :
 Proof.
   apply algo_dom_mutual.
   - move => a b h.
-    move => h0 h1.
-    have {}h1 : check_equal_r a b h by hauto l:on rew:db:check_equal.
+    move => h0.
+    rewrite check_equal_abs_abs.
     constructor. tauto.
   - move => a u i h0 ih h.
     apply CE_AbsNeu => //.
@@ -116,10 +130,11 @@ Proof.
   - move => a u i h ih h0.
     apply CE_NeuAbs=>//.
     apply ih.
-    case : u i h ih h0 => //=.
+    by rewrite check_equal_neu_abs in h0.
   - move => a0 a1 b0 b1 a ha h.
-    move => h0 h1. constructor. apply ha. hauto lb:on rew:db:check_equal.
-    apply h0. hauto lb:on rew:db:check_equal.
+    move => h0.
+    rewrite check_equal_pair_pair. move /andP => [h1 h2].
+    sauto lq:on.
   - move => a0 a1 u neu h ih h' ih' he.
     rewrite check_equal_pair_neu in he.
     apply CE_PairNeu => //; hauto lqb:on.

From b9b6899764241ce6ce969c4e5e1b5c257a45c0b4 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 4 Mar 2025 00:39:59 -0500
Subject: [PATCH 16/24] Half way done with check_equal_complete

---
 theories/executable_correct.v | 27 +++++++++++++++++++++++++++
 1 file changed, 27 insertions(+)

diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index d520857..54566c8 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -107,6 +107,8 @@ Proof.
     sfirstorder use:hne_no_hred, hf_no_hred.
 Qed.
 
+#[export]Hint Rewrite check_equal_abs_abs check_equal_abs_neu check_equal_neu_abs check_equal_pair_pair check_equal_pair_neu check_equal_neu_pair check_equal_bind_bind check_equal_hredl check_equal_hredr check_equal_nfnf : ce_prop.
+
 Lemma coqeq_neuneu u0 u1 :
   ishne u0 -> ishne u1 -> u0 ↔ u1 -> u0 ∼ u1.
 Proof.
@@ -175,3 +177,28 @@ Proof.
   - move => a b b' hu r a0 ha hb. rewrite check_equal_hredr in hb.
     sauto lq:on rew:off.
 Qed.
+
+Lemma check_equal_complete :
+  (forall a b (h : algo_dom a b), ~ check_equal a b h -> ~ a ↔ b) /\
+  (forall a b (h : algo_dom_r a b), ~ check_equal_r a b h -> ~ a ⇔ b).
+Proof.
+  apply algo_dom_mutual.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - move => i j.
+    move => h0 h1.
+    have ? : i = j by sauto lq:on. subst.
+    simp check_equal in h0.
+    set x := (nat_eqdec j j).
+    destruct x as [].
+    sauto q:on.
+    sfirstorder.
+  - admit.
+  - sfirstorder.
+  - best.

From dcd4465310d86d6de785f8cbff6a0e9d6cdf28c1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 21:47:57 -0500
Subject: [PATCH 17/24] Finish the proof of completeness for the algorithm

---
 theories/executable_correct.v | 74 +++++++++++++++++++++++++++++++++--
 1 file changed, 71 insertions(+), 3 deletions(-)

diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 54566c8..27092b1 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -62,6 +62,14 @@ Proof.
   sfirstorder use:hred_complete, hred_sound.
 Qed.
 
+Lemma coqeqr_no_hred a b : a ∼ b -> HRed.nf a /\ HRed.nf b.
+Proof. induction 1; sauto lq:on unfold:HRed.nf. Qed.
+
+Lemma coqeq_no_hred a b : a ↔ b -> HRed.nf a /\ HRed.nf b.
+Proof. induction 1;
+         qauto inv:HRed.R use:coqeqr_no_hred, hne_no_hred unfold:HRed.nf.
+Qed.
+
 Lemma check_equal_nfnf a b dom : check_equal_r a b (A_NfNf a b dom) = check_equal a b dom.
 Proof.
   have [h0 h1] : (ishf a \/ ishne a) /\ (ishf b \/ ishne b) by hauto l:on use:algo_dom_hf_hne.
@@ -178,6 +186,12 @@ Proof.
     sauto lq:on rew:off.
 Qed.
 
+Lemma hreds_nf_refl a b  :
+  HRed.nf a ->
+  rtc HRed.R a b ->
+  a = b.
+Proof. inversion 2; sfirstorder. Qed.
+
 Lemma check_equal_complete :
   (forall a b (h : algo_dom a b), ~ check_equal a b h -> ~ a ↔ b) /\
   (forall a b (h : algo_dom_r a b), ~ check_equal_r a b h -> ~ a ⇔ b).
@@ -199,6 +213,60 @@ Proof.
     destruct x as [].
     sauto q:on.
     sfirstorder.
-  - admit.
-  - sfirstorder.
-  - best.
+  - move => p0 p1 A0 A1 B0 B1 h0 ih0 h1 ih1.
+    rewrite check_equal_bind_bind.
+    move /negP.
+    move /nandP.
+    case. move /nandP.
+    case. move => h. have : p0 <> p1  by sauto lqb:on.
+    clear. move => ?. hauto lq:on rew:off inv:CoqEq, CoqEq_Neu.
+    hauto qb:on inv:CoqEq,CoqEq_Neu.
+    hauto qb:on inv:CoqEq,CoqEq_Neu.
+  - simp check_equal. done.
+  - move => i j. simp check_equal.
+    case : nat_eqdec => //=. sauto lq:on.
+  - move => p0 p1 u0 u1 neu0 neu1 h ih.
+    rewrite check_equal_proj_proj.
+    move /negP /nandP. case.
+    case : PTag_eqdec => //=. sauto lq:on.
+    sauto lqb:on.
+  - move => u0 u1 a0 a1 hu0 hu1 h0 ih0 h1 ih1.
+    rewrite check_equal_app_app.
+    move /negP /nandP. sauto b:on inv:CoqEq, CoqEq_Neu.
+  - move => P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 domP ihP domu ihu domb ihb domc ihc.
+    rewrite check_equal_ind_ind.
+    move => + h.
+    inversion h; subst.
+    inversion H; subst.
+    move /negP /nandP.
+    case. move/nandP.
+    case. move/nandP.
+    case. qauto b:on inv:CoqEq, CoqEq_Neu.
+    sauto lqb:on inv:CoqEq, CoqEq_Neu.
+    sauto lqb:on inv:CoqEq, CoqEq_Neu.
+    sauto lqb:on inv:CoqEq, CoqEq_Neu.
+  - move => a b h ih.
+    rewrite check_equal_nfnf.
+    move : ih => /[apply].
+    move => + h0.
+    move /algo_dom_hf_hne in h.
+    inversion h0; subst.
+    have {h} [? ?] : HRed.nf a /\ HRed.nf b by sfirstorder use:hf_no_hred, hne_no_hred.
+    hauto l:on use:hreds_nf_refl.
+  - move => a a' b h dom.
+    simp ce_prop => /[apply].
+    move => + h1. inversion h1; subst.
+    inversion H; subst.
+    + sfirstorder use:coqeq_no_hred unfold:HRed.nf.
+    + have ? : y = a' by eauto using hred_deter. subst.
+      sauto lq:on.
+  - move => a b b' u hr dom ihdom.
+    rewrite check_equal_hredr.
+    move => + h. inversion h; subst.
+    have {}u : HRed.nf a by sfirstorder use:hne_no_hred, hf_no_hred.
+    move => {}/ihdom.
+    inversion H0; subst.
+    + have ? : HRed.nf b'0 by hauto l:on use:coqeq_no_hred.
+      sfirstorder unfold:HRed.nf.
+    + sauto lq:on use:hred_deter.
+Qed.

From 5087b8c6ce225d7e3f782769dd1cdf13e62639f4 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 22:20:12 -0500
Subject: [PATCH 18/24] Add new definition of eq_view

---
 theories/executable.v | 33 ++++++++++++++++++++++++---------
 1 file changed, 24 insertions(+), 9 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index 5de3355..3ffc695 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -30,20 +30,20 @@ Inductive eq_view : PTm -> PTm -> Type :=
 | V_AbsAbs a b :
   eq_view (PAbs a) (PAbs b)
 | V_AbsNeu a b :
-  ~~ isabs b ->
+  ~~ ishf b ->
   eq_view (PAbs a) b
 | V_NeuAbs a b :
-  ~~ isabs a ->
+  ~~ ishf a ->
   eq_view a (PAbs b)
 | V_VarVar i j :
   eq_view (VarPTm i) (VarPTm j)
 | V_PairPair a0 b0 a1 b1 :
   eq_view (PPair a0 b0) (PPair a1 b1)
 | V_PairNeu a0 b0 u :
-  ~~ ispair u ->
+  ~~ ishf u ->
   eq_view (PPair a0 b0) u
 | V_NeuPair u a1 b1 :
-  ~~ ispair u ->
+  ~~ ishf u ->
   eq_view u (PPair a1 b1)
 | V_ZeroZero :
   eq_view PZero PZero
@@ -62,16 +62,31 @@ Inductive eq_view : PTm -> PTm -> Type :=
 | V_UnivUniv i j :
   eq_view (PUniv i) (PUniv j)
 | V_Others a b :
+  tm_conf a b ->
   eq_view a b.
 
 Equations tm_to_eq_view (a b : PTm) : eq_view a b :=
   tm_to_eq_view (PAbs a) (PAbs b) := V_AbsAbs a b;
-  tm_to_eq_view (PAbs a) b := V_AbsNeu a b _;
-  tm_to_eq_view a (PAbs b) := V_NeuAbs a b _;
+  tm_to_eq_view (PAbs a) (PApp b0 b1) := V_AbsNeu a (PApp b0 b1) _;
+  tm_to_eq_view (PAbs a) (VarPTm i) := V_AbsNeu a (VarPTm i) _;
+  tm_to_eq_view (PAbs a) (PProj p b) := V_AbsNeu a (PProj p b) _;
+  tm_to_eq_view (PAbs a) (PInd P u b c) := V_AbsNeu a (PInd P u b c) _;
+  tm_to_eq_view (VarPTm i) (PAbs a) := V_NeuAbs (VarPTm i) a _;
+  tm_to_eq_view (PApp b0 b1) (PAbs b) := V_NeuAbs (PApp b0 b1) b _;
+  tm_to_eq_view (PProj p b) (PAbs a) := V_NeuAbs (PProj p b) a _;
+  tm_to_eq_view (PInd P u b c) (PAbs a) := V_NeuAbs (PInd P u b c) a _;
   tm_to_eq_view (VarPTm i) (VarPTm j) := V_VarVar i j;
   tm_to_eq_view (PPair a0 b0) (PPair a1 b1) := V_PairPair a0 b0 a1 b1;
-  tm_to_eq_view (PPair a0 b0) u := V_PairNeu a0 b0 u _;
-  tm_to_eq_view u (PPair a1 b1) := V_NeuPair u a1 b1 _;
+  (* tm_to_eq_view (PPair a0 b0) u := V_PairNeu a0 b0 u _; *)
+  tm_to_eq_view (PPair a0 b0) (VarPTm i) := V_PairNeu a0 b0 (VarPTm i) _;
+  tm_to_eq_view (PPair a0 b0) (PApp c0 c1) := V_PairNeu a0 b0 (PApp c0 c1) _;
+  tm_to_eq_view (PPair a0 b0) (PProj p c) := V_PairNeu a0 b0 (PProj p c) _;
+  tm_to_eq_view (PPair a0 b0) (PInd P t0 t1 t2) := V_PairNeu a0 b0 (PInd P t0 t1 t2) _;
+  (* tm_to_eq_view u (PPair a1 b1) := V_NeuPair u a1 b1 _; *)
+  tm_to_eq_view (VarPTm i) (PPair a1 b1) := V_NeuPair (VarPTm i) a1 b1 _;
+  tm_to_eq_view (PApp t0 t1) (PPair a1 b1) := V_NeuPair (PApp t0 t1) a1 b1 _;
+  tm_to_eq_view (PProj t0 t1) (PPair a1 b1) := V_NeuPair (PProj t0 t1) a1 b1 _;
+  tm_to_eq_view (PInd t0 t1 t2 t3) (PPair a1 b1) := V_NeuPair (PInd t0 t1 t2 t3) a1 b1 _;
   tm_to_eq_view PZero PZero := V_ZeroZero;
   tm_to_eq_view (PSuc a) (PSuc b) := V_SucSuc a b;
   tm_to_eq_view (PApp a0 b0) (PApp a1 b1) := V_AppApp a0 b0 a1 b1;
@@ -80,7 +95,7 @@ Equations tm_to_eq_view (a b : PTm) : eq_view a b :=
   tm_to_eq_view PNat PNat := V_NatNat;
   tm_to_eq_view (PUniv i) (PUniv j) := V_UnivUniv i j;
   tm_to_eq_view (PBind p0 A0 B0)  (PBind p1 A1 B1) := V_BindBind p0 A0 B0 p1 A1 B1;
-  tm_to_eq_view a b := V_Others a b.
+  tm_to_eq_view a b := V_Others a b _.
 
 Inductive algo_dom : PTm -> PTm -> Prop :=
 | A_AbsAbs a b :

From a23be7f9b50da2d1c621628896be7e435a75d642 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 22:28:18 -0500
Subject: [PATCH 19/24] Simplify the definition of algo_dom

---
 theories/executable.v | 27 +++++++++++++++------------
 1 file changed, 15 insertions(+), 12 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index 3ffc695..159f669 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -183,6 +183,10 @@ Inductive algo_dom : PTm -> PTm -> Prop :=
   algo_dom_r c0 c1 ->
   algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
 
+(* | A_Conf a b : *)
+(*   tm_conf a b -> *)
+(*   algo_dom a b *)
+
 with algo_dom_r : PTm  -> PTm -> Prop :=
 | A_NfNf a b :
   algo_dom a b ->
@@ -195,19 +199,12 @@ with algo_dom_r : PTm  -> PTm -> Prop :=
   algo_dom_r a b
 
 | A_HRedR a b b'  :
-  ishne a \/ ishf a ->
+  HRed.nf a ->
   HRed.R b b' ->
   algo_dom_r a b' ->
   (* ----------------------- *)
   algo_dom_r a b.
 
-Lemma algo_dom_hf_hne (a b : PTm) :
-  algo_dom a b ->
-  (ishf a \/ ishne a) /\ (ishf b \/ ishne b).
-Proof.
-  induction 1 =>//=; hauto lq:on.
-Qed.
-
 Lemma hf_no_hred (a b : PTm) :
   ishf a ->
   HRed.R a b ->
@@ -220,6 +217,13 @@ Lemma hne_no_hred (a b : PTm) :
   False.
 Proof. elim : a b => //=; hauto l:on inv:HRed.R. Qed.
 
+Lemma algo_dom_no_hred (a b : PTm) :
+  algo_dom a b ->
+  HRed.nf a /\ HRed.nf b.
+Proof.
+  induction 1 =>//=; try hauto inv:HRed.R use:hne_no_hred, hf_no_hred lq:on unfold:HRed.nf.
+Qed.
+
 Derive Signature for algo_dom algo_dom_r.
 
 Fixpoint hred (a : PTm) : option (PTm) :=
@@ -342,9 +346,7 @@ Defined.
 Next Obligation.
   intros.
   destruct h.
-  exfalso. apply algo_dom_hf_hne in H0.
-  destruct H0 as [h0 h1].
-  sfirstorder use:hf_no_hred, hne_no_hred.
+  exfalso. sfirstorder use: algo_dom_no_hred.
   exfalso. sfirstorder.
   assert (  b' = b'0)by eauto using hred_deter. subst.
   apply h.
@@ -353,7 +355,8 @@ Defined.
 Next Obligation.
   simpl. intros.
   inversion h; subst =>//=.
-  move {h}. hauto lq:on use:algo_dom_hf_hne, hf_no_hred, hne_no_hred, hred_sound.
+  exfalso. sfirstorder use:algo_dom_no_hred.
+  move {h}.
   assert (a' = a'0) by eauto using hred_deter, hred_sound. by subst.
   exfalso. sfirstorder use:hne_no_hred, hf_no_hred.
 Defined.

From 68cc482479519553e335f63f81e54d2ac106bbc6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 22:30:21 -0500
Subject: [PATCH 20/24] Fix the correctness proof

---
 theories/executable_correct.v | 5 ++---
 1 file changed, 2 insertions(+), 3 deletions(-)

diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 27092b1..1464f44 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -72,7 +72,7 @@ Qed.
 
 Lemma check_equal_nfnf a b dom : check_equal_r a b (A_NfNf a b dom) = check_equal a b dom.
 Proof.
-  have [h0 h1] : (ishf a \/ ishne a) /\ (ishf b \/ ishne b) by hauto l:on use:algo_dom_hf_hne.
+  have [h0 h1] : HRed.nf a /\ HRed.nf b by hauto l:on use:algo_dom_no_hred.
   have [h3 h4] : hred a = None /\ hred b = None by sfirstorder use:hf_no_hred, hne_no_hred, hred_none.
   simp check_equal.
   destruct (fancy_hred a).
@@ -249,9 +249,8 @@ Proof.
     rewrite check_equal_nfnf.
     move : ih => /[apply].
     move => + h0.
-    move /algo_dom_hf_hne in h.
+    have {h} [? ?] : HRed.nf a /\ HRed.nf b by sfirstorder use:algo_dom_no_hred.
     inversion h0; subst.
-    have {h} [? ?] : HRed.nf a /\ HRed.nf b by sfirstorder use:hf_no_hred, hne_no_hred.
     hauto l:on use:hreds_nf_refl.
   - move => a a' b h dom.
     simp ce_prop => /[apply].

From c05bd100166626c6c92c15e205a3db621d26dea5 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 22:43:30 -0500
Subject: [PATCH 21/24] Turn off auto equations generation because it produces
 poor lemmas

---
 theories/executable.v         | 110 ++++++++++++++++++++++++++++++++--
 theories/executable_correct.v |  98 ------------------------------
 2 files changed, 106 insertions(+), 102 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index 159f669..fa3f51a 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -1,5 +1,6 @@
 From Equations Require Import Equations.
 Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
+Require Import Logic.PropExtensionality (propositional_extensionality).
 Require Import ssreflect ssrbool.
 Import Logic (inspect).
 
@@ -297,7 +298,7 @@ Ltac solve_check_equal :=
   | _ => idtac
   end].
 
-Equations check_equal (a b : PTm) (h : algo_dom a b) :
+#[derive(equations=no)]Equations check_equal (a b : PTm) (h : algo_dom a b) :
   bool by struct h :=
   check_equal a b h with tm_to_eq_view a b :=
   check_equal _ _ h (V_VarVar i j) := nat_eqdec i j;
@@ -364,6 +365,107 @@ Defined.
 
 
 
-Next Obligation.
-  qauto inv:algo_dom, algo_dom_r.
-Defined.
+(* Next Obligation. *)
+(*   qauto inv:algo_dom, algo_dom_r. *)
+(* Defined. *)
+
+Lemma check_equal_abs_abs a b h : check_equal (PAbs a) (PAbs b) (A_AbsAbs a b h) = check_equal_r a b h.
+Proof. reflexivity. Qed.
+
+Lemma check_equal_abs_neu a u neu h : check_equal (PAbs a) u (A_AbsNeu a u neu h) = check_equal_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) h.
+Proof. case : u neu h => //=.  Qed.
+
+Lemma check_equal_neu_abs a u neu h : check_equal u (PAbs a) (A_NeuAbs a u neu h) = check_equal_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a h.
+Proof. case : u neu h => //=.  Qed.
+
+Lemma check_equal_pair_pair a0 b0 a1 b1 a h :
+  check_equal (PPair a0 b0) (PPair a1 b1) (A_PairPair a0 a1 b0 b1 a h) =
+    check_equal_r a0 a1 a && check_equal_r b0 b1 h.
+Proof. reflexivity. Qed.
+
+Lemma check_equal_pair_neu a0 a1 u neu h h'
+  : check_equal (PPair a0 a1) u (A_PairNeu a0 a1 u neu h h') = check_equal_r a0 (PProj PL u) h && check_equal_r a1 (PProj PR u) h'.
+Proof.
+  case : u neu h h' => //=.
+Qed.
+
+Lemma check_equal_neu_pair a0 a1 u neu h h'
+  : check_equal u (PPair a0 a1) (A_NeuPair a0 a1 u neu h h') = check_equal_r  (PProj PL u) a0 h && check_equal_r (PProj PR u) a1 h'.
+Proof.
+  case : u neu h h' => //=.
+Qed.
+
+Lemma check_equal_bind_bind p0 A0 B0 p1 A1 B1 h0 h1 :
+  check_equal (PBind p0 A0 B0) (PBind p1 A1 B1) (A_BindCong p0 p1 A0 A1 B0 B1 h0 h1) =
+    BTag_eqdec p0 p1 &&  check_equal_r A0 A1 h0 && check_equal_r  B0 B1 h1.
+Proof. reflexivity. Qed.
+
+Lemma check_equal_proj_proj p0 u0 p1 u1 neu0 neu1 h :
+  check_equal (PProj p0 u0) (PProj p1 u1) (A_ProjCong p0 p1 u0 u1 neu0 neu1 h) =
+    PTag_eqdec p0 p1 && check_equal u0 u1 h.
+Proof. reflexivity. Qed.
+
+Lemma check_equal_app_app u0 a0 u1 a1 hu0 hu1 hdom hdom' :
+  check_equal (PApp u0 a0) (PApp u1 a1) (A_AppCong u0 u1 a0 a1 hu0 hu1 hdom hdom') =
+    check_equal u0 u1 hdom && check_equal_r a0 a1 hdom'.
+Proof. reflexivity. Qed.
+
+Lemma check_equal_ind_ind P0 u0 b0 c0 P1 u1 b1 c1 neu0 neu1 domP domu domb domc :
+  check_equal (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
+    (A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 domP domu domb domc) =
+    check_equal_r P0 P1 domP && check_equal u0 u1 domu && check_equal_r b0 b1 domb && check_equal_r c0 c1 domc.
+Proof. reflexivity. Qed.
+
+Lemma hred_none a : HRed.nf a -> hred a = None.
+Proof.
+  destruct (hred a) eqn:eq;
+  sfirstorder use:hred_complete, hred_sound.
+Qed.
+
+Lemma check_equal_nfnf a b dom : check_equal_r a b (A_NfNf a b dom) = check_equal a b dom.
+Proof.
+  have [h0 h1] : HRed.nf a /\ HRed.nf b by hauto l:on use:algo_dom_no_hred.
+  have [h3 h4] : hred a = None /\ hred b = None by sfirstorder use:hf_no_hred, hne_no_hred, hred_none.
+  simpl.
+  rewrite /check_equal_r_functional.
+  destruct (fancy_hred a).
+  simpl.
+  destruct (fancy_hred b).
+  reflexivity.
+  exfalso. hauto l:on use:hred_complete.
+  exfalso. hauto l:on use:hred_complete.
+Qed.
+
+Lemma check_equal_hredl a b a' ha doma :
+  check_equal_r a b (A_HRedL a a' b ha doma) = check_equal_r a' b doma.
+Proof.
+  simpl.
+  rewrite /check_equal_r_functional.
+  destruct (fancy_hred a).
+  - hauto q:on unfold:HRed.nf.
+  - destruct s as [x ?].
+    rewrite /check_equal_r_functional.
+    have ? : x = a' by eauto using hred_deter. subst.
+    simpl.
+    f_equal.
+    apply PropExtensionality.proof_irrelevance.
+Qed.
+
+Lemma check_equal_hredr a b b' hu r a0 :
+  check_equal_r a b (A_HRedR a b b' hu r a0) =
+    check_equal_r a b' a0.
+Proof.
+  simpl. rewrite /check_equal_r_functional.
+  destruct (fancy_hred a).
+  - simpl.
+    destruct (fancy_hred b) as [|[b'' hb']].
+    + hauto lq:on unfold:HRed.nf.
+    + simpl.
+      have ? : (b'' = b') by eauto using hred_deter. subst.
+      f_equal.
+      apply PropExtensionality.proof_irrelevance.
+  - exfalso.
+    sfirstorder use:hne_no_hred, hf_no_hred.
+Qed.
+
+#[export]Hint Rewrite check_equal_abs_abs check_equal_abs_neu check_equal_neu_abs check_equal_pair_pair check_equal_pair_neu check_equal_neu_pair check_equal_bind_bind check_equal_hredl check_equal_hredr check_equal_nfnf : ce_prop.
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 1464f44..1a28e23 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -9,58 +9,6 @@ Scheme algo_ind := Induction for algo_dom Sort Prop
 
 Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
 
-Lemma check_equal_abs_abs a b h : check_equal (PAbs a) (PAbs b) (A_AbsAbs a b h) = check_equal_r a b h.
-Proof. hauto l:on rew:db:check_equal. Qed.
-
-Lemma check_equal_abs_neu a u neu h : check_equal (PAbs a) u (A_AbsNeu a u neu h) = check_equal_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) h.
-Proof. case : u neu h => //=.  Qed.
-
-Lemma check_equal_neu_abs a u neu h : check_equal u (PAbs a) (A_NeuAbs a u neu h) = check_equal_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a h.
-Proof. case : u neu h => //=.  Qed.
-
-Lemma check_equal_pair_pair a0 b0 a1 b1 a h :
-  check_equal (PPair a0 b0) (PPair a1 b1) (A_PairPair a0 a1 b0 b1 a h) =
-    check_equal_r a0 a1 a && check_equal_r b0 b1 h.
-Proof. hauto l:on rew:db:check_equal. Qed.
-
-Lemma check_equal_pair_neu a0 a1 u neu h h'
-  : check_equal (PPair a0 a1) u (A_PairNeu a0 a1 u neu h h') = check_equal_r a0 (PProj PL u) h && check_equal_r a1 (PProj PR u) h'.
-Proof.
-  case : u neu h h' => //=; simp check_equal tm_to_eq_view.
-Qed.
-
-Lemma check_equal_neu_pair a0 a1 u neu h h'
-  : check_equal u (PPair a0 a1) (A_NeuPair a0 a1 u neu h h') = check_equal_r  (PProj PL u) a0 h && check_equal_r (PProj PR u) a1 h'.
-Proof.
-  case : u neu h h' => //=; simp check_equal tm_to_eq_view.
-Qed.
-
-Lemma check_equal_bind_bind p0 A0 B0 p1 A1 B1 h0 h1 :
-  check_equal (PBind p0 A0 B0) (PBind p1 A1 B1) (A_BindCong p0 p1 A0 A1 B0 B1 h0 h1) =
-    BTag_eqdec p0 p1 &&  check_equal_r A0 A1 h0 && check_equal_r  B0 B1 h1.
-Proof. hauto lq:on. Qed.
-
-Lemma check_equal_proj_proj p0 u0 p1 u1 neu0 neu1 h :
-  check_equal (PProj p0 u0) (PProj p1 u1) (A_ProjCong p0 p1 u0 u1 neu0 neu1 h) =
-    PTag_eqdec p0 p1 && check_equal u0 u1 h.
-Proof. hauto lq:on. Qed.
-
-Lemma check_equal_app_app u0 a0 u1 a1 hu0 hu1 hdom hdom' :
-  check_equal (PApp u0 a0) (PApp u1 a1) (A_AppCong u0 u1 a0 a1 hu0 hu1 hdom hdom') =
-    check_equal u0 u1 hdom && check_equal_r a0 a1 hdom'.
-Proof. hauto lq:on. Qed.
-
-Lemma check_equal_ind_ind P0 u0 b0 c0 P1 u1 b1 c1 neu0 neu1 domP domu domb domc :
-  check_equal (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
-    (A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 domP domu domb domc) =
-    check_equal_r P0 P1 domP && check_equal u0 u1 domu && check_equal_r b0 b1 domb && check_equal_r c0 c1 domc.
-Proof. hauto lq:on. Qed.
-
-Lemma hred_none a : HRed.nf a -> hred a = None.
-Proof.
-  destruct (hred a) eqn:eq;
-  sfirstorder use:hred_complete, hred_sound.
-Qed.
 
 Lemma coqeqr_no_hred a b : a ∼ b -> HRed.nf a /\ HRed.nf b.
 Proof. induction 1; sauto lq:on unfold:HRed.nf. Qed.
@@ -70,52 +18,6 @@ Proof. induction 1;
          qauto inv:HRed.R use:coqeqr_no_hred, hne_no_hred unfold:HRed.nf.
 Qed.
 
-Lemma check_equal_nfnf a b dom : check_equal_r a b (A_NfNf a b dom) = check_equal a b dom.
-Proof.
-  have [h0 h1] : HRed.nf a /\ HRed.nf b by hauto l:on use:algo_dom_no_hred.
-  have [h3 h4] : hred a = None /\ hred b = None by sfirstorder use:hf_no_hred, hne_no_hred, hred_none.
-  simp check_equal.
-  destruct (fancy_hred a).
-  simp check_equal.
-  destruct (fancy_hred b).
-  simp check_equal. hauto lq:on.
-  exfalso. hauto l:on use:hred_complete.
-  exfalso. hauto l:on use:hred_complete.
-Qed.
-
-Lemma check_equal_hredl a b a' ha doma :
-  check_equal_r a b (A_HRedL a a' b ha doma) = check_equal_r a' b doma.
-Proof.
-  simp check_equal.
-  destruct (fancy_hred a).
-  - hauto q:on unfold:HRed.nf.
-  - simp check_equal.
-    destruct s as [x ?]. have ? : x = a' by eauto using hred_deter. subst.
-    simpl.
-    simp check_equal.
-    f_equal.
-    apply PropExtensionality.proof_irrelevance.
-Qed.
-
-Lemma check_equal_hredr a b b' hu r a0 :
-  check_equal_r a b (A_HRedR a b b' hu r a0) =
-    check_equal_r a b' a0.
-Proof.
-  simp check_equal.
-  destruct (fancy_hred a).
-  - rewrite check_equal_r_clause_1_equation_1.
-    destruct (fancy_hred b) as [|[b'' hb']].
-    + hauto lq:on unfold:HRed.nf.
-    + have ? : (b'' = b') by eauto using hred_deter. subst.
-      rewrite check_equal_r_clause_1_equation_1.  simpl.
-      simp check_equal. destruct (fancy_hred a). simp check_equal.
-      f_equal; apply PropExtensionality.proof_irrelevance.
-      simp check_equal. exfalso. sfirstorder use:hne_no_hred, hf_no_hred.
-  - simp check_equal. exfalso.
-    sfirstorder use:hne_no_hred, hf_no_hred.
-Qed.
-
-#[export]Hint Rewrite check_equal_abs_abs check_equal_abs_neu check_equal_neu_abs check_equal_pair_pair check_equal_pair_neu check_equal_neu_pair check_equal_bind_bind check_equal_hredl check_equal_hredr check_equal_nfnf : ce_prop.
 
 Lemma coqeq_neuneu u0 u1 :
   ishne u0 -> ishne u1 -> u0 ↔ u1 -> u0 ∼ u1.

From c1ff0ae145017f0f7ef0b4c9edea096a3eb9cbee Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 23:22:41 -0500
Subject: [PATCH 22/24] Add check_equal_conf case

---
 theories/executable.v | 51 +++++++++++++++++++++++++++----------------
 1 file changed, 32 insertions(+), 19 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index fa3f51a..f7991b0 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -3,22 +3,26 @@ Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 Require Import Logic.PropExtensionality (propositional_extensionality).
 Require Import ssreflect ssrbool.
 Import Logic (inspect).
+From Ltac2 Require Import Ltac2.
+Import Ltac2.Constr.
+Set Default Proof Mode "Classic".
+
 
 Require Import ssreflect ssrbool.
 From Hammer Require Import Tactics.
 
 Definition tm_nonconf (a b : PTm) : bool :=
   match a, b with
-  | PAbs _, _ => ishne b || isabs b
-  | _, PAbs _ => ishne a
+  | PAbs _, _ => (~~ ishf b) || isabs b
+  | _, PAbs _ => ~~ ishf a
   | VarPTm _, VarPTm _ => true
-  | PPair _ _, _ => ishne b || ispair b
-  | _, PPair _ _ => ishne a
+  | PPair _ _, _ => (~~ ishf b) || ispair b
+  | _, PPair _ _ => ~~ ishf a
   | PZero, PZero => true
   | PSuc _, PSuc _ => true
-  | PApp _ _, PApp _ _ => ishne a && ishne b
-  | PProj _ _, PProj _ _ => ishne a && ishne b
-  | PInd _ _ _ _, PInd _ _ _ _ => ishne a && ishne b
+  | PApp _ _, PApp _ _ => (~~ ishf a) && (~~ ishf b)
+  | PProj _ _, PProj _ _ => (~~ ishf a) && (~~ ishf b)
+  | PInd _ _ _ _, PInd _ _ _ _ => (~~ ishf a) && (~~ ishf b)
   | PNat, PNat => true
   | PUniv _, PUniv _ => true
   | PBind _ _ _, PBind _ _ _ => true
@@ -184,9 +188,11 @@ Inductive algo_dom : PTm -> PTm -> Prop :=
   algo_dom_r c0 c1 ->
   algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
 
-(* | A_Conf a b : *)
-(*   tm_conf a b -> *)
-(*   algo_dom a b *)
+| A_Conf a b :
+  HRed.nf a ->
+  HRed.nf b ->
+  tm_conf a b ->
+  algo_dom a b
 
 with algo_dom_r : PTm  -> PTm -> Prop :=
 | A_NfNf a b :
@@ -283,11 +289,19 @@ Definition fancy_hred (a : PTm) : HRed.nf a + {b | HRed.R a b}.
   left. move => b /hred_complete. congruence.
 Defined.
 
+Ltac2 destruct_algo () :=
+  lazy_match! goal with
+  | [h : algo_dom ?a ?b |- _ ] =>
+      if is_var a then destruct $a; ltac1:(done)  else
+        (if is_var b then destruct $b; ltac1:(done) else ())
+  end.
+
+
 Ltac check_equal_triv :=
   intros;subst;
   lazymatch goal with
   (* | [h : algo_dom (VarPTm _) (PAbs _) |- _] => idtac *)
-  | [h : algo_dom _ _ |- _] => try (inversion h; by subst)
+  | [h : algo_dom _ _ |- _] => try (inversion h; subst => //=; ltac2:(Control.enter destruct_algo))
   | _ => idtac
   end.
 
@@ -335,7 +349,8 @@ Ltac solve_check_equal :=
       check_equal_r a b h (inr a') := check_equal_r (proj1_sig a') b _;
       check_equal_r a b h (inl h')  with (fancy_hred b) :=
         | inr b' := check_equal_r a (proj1_sig b') _;
-        | inl h'' := check_equal a b _.
+                 | inl h'' := check_equal a b _.
+
 
 Next Obligation.
   intros.
@@ -363,12 +378,6 @@ Next Obligation.
 Defined.
 
 
-
-
-(* Next Obligation. *)
-(*   qauto inv:algo_dom, algo_dom_r. *)
-(* Defined. *)
-
 Lemma check_equal_abs_abs a b h : check_equal (PAbs a) (PAbs b) (A_AbsAbs a b h) = check_equal_r a b h.
 Proof. reflexivity. Qed.
 
@@ -468,4 +477,8 @@ Proof.
     sfirstorder use:hne_no_hred, hf_no_hred.
 Qed.
 
-#[export]Hint Rewrite check_equal_abs_abs check_equal_abs_neu check_equal_neu_abs check_equal_pair_pair check_equal_pair_neu check_equal_neu_pair check_equal_bind_bind check_equal_hredl check_equal_hredr check_equal_nfnf : ce_prop.
+Lemma check_equal_conf a b nfa nfb nfab :
+  check_equal a b (A_Conf a b nfa nfb nfab) = false.
+Proof. destruct a; destruct b => //=. Qed.
+
+#[export]Hint Rewrite check_equal_abs_abs check_equal_abs_neu check_equal_neu_abs check_equal_pair_pair check_equal_pair_neu check_equal_neu_pair check_equal_bind_bind check_equal_hredl check_equal_hredr check_equal_nfnf check_equal_conf : ce_prop.

From c4f01d7dfc4fd4fa94b5c8760ad3c86741471fae Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 4 Mar 2025 23:48:42 -0500
Subject: [PATCH 23/24] Update the correctness proof of the computable function

---
 theories/executable.v         |  4 ++++
 theories/executable_correct.v | 44 ++++++++++++++++++++---------------
 2 files changed, 29 insertions(+), 19 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index f7991b0..0586a41 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -477,6 +477,10 @@ Proof.
     sfirstorder use:hne_no_hred, hf_no_hred.
 Qed.
 
+Lemma check_equal_univ i j :
+  check_equal (PUniv i) (PUniv j) (A_UnivCong i j) = nat_eqdec i j.
+Proof. reflexivity. Qed.
+
 Lemma check_equal_conf a b nfa nfb nfab :
   check_equal a b (A_Conf a b nfa nfb nfab) = false.
 Proof. destruct a; destruct b => //=. Qed.
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 1a28e23..0720d03 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -37,8 +37,7 @@ Proof.
   - move => a u i h0 ih h.
     apply CE_AbsNeu => //.
     apply : ih. simp check_equal tm_to_eq_view in h.
-    have h1 : check_equal (PAbs a) u (A_AbsNeu a u i h0) = check_equal_r a (PApp (ren_PTm shift u) (VarPTm var_zero)) h0 by  clear; case : u i h0 => //=.
-    hauto lq:on.
+    by rewrite check_equal_abs_neu in h.
   - move => a u i h ih h0.
     apply CE_NeuAbs=>//.
     apply ih.
@@ -79,6 +78,7 @@ Proof.
     rewrite check_equal_ind_ind.
     move /andP => [/andP [/andP [h0 h1] h2 ] h3].
     sauto lq:on use:coqeq_neuneu.
+  - move => a b n n0 i. by rewrite check_equal_conf.
   - move => a b dom h ih. apply : CE_HRed; eauto using rtc_refl.
     rewrite check_equal_nfnf in ih.
     tauto.
@@ -94,27 +94,25 @@ Lemma hreds_nf_refl a b  :
   a = b.
 Proof. inversion 2; sfirstorder. Qed.
 
+Ltac ce_solv := move => *; simp ce_prop in *; hauto qb:on rew:off inv:CoqEq, CoqEq_Neu.
+
 Lemma check_equal_complete :
   (forall a b (h : algo_dom a b), ~ check_equal a b h -> ~ a ↔ b) /\
   (forall a b (h : algo_dom_r a b), ~ check_equal_r a b h -> ~ a ⇔ b).
 Proof.
   apply algo_dom_mutual.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
-  - hauto q:on inv:CoqEq, CoqEq_Neu b:on rew:db:ce_prop.
+  - ce_solv.
+  - ce_solv.
+  - ce_solv.
+  - ce_solv.
+  - ce_solv.
+  - ce_solv.
+  - ce_solv.
+  - ce_solv.
   - move => i j.
-    move => h0 h1.
-    have ? : i = j by sauto lq:on. subst.
-    simp check_equal in h0.
-    set x := (nat_eqdec j j).
-    destruct x as [].
-    sauto q:on.
-    sfirstorder.
+    rewrite check_equal_univ.
+    case : nat_eqdec => //=.
+    ce_solv.
   - move => p0 p1 A0 A1 B0 B1 h0 ih0 h1 ih1.
     rewrite check_equal_bind_bind.
     move /negP.
@@ -125,8 +123,8 @@ Proof.
     hauto qb:on inv:CoqEq,CoqEq_Neu.
     hauto qb:on inv:CoqEq,CoqEq_Neu.
   - simp check_equal. done.
-  - move => i j. simp check_equal.
-    case : nat_eqdec => //=. sauto lq:on.
+  - move => i j. move => h. have {h} : ~ nat_eqdec i j by done.
+    case : nat_eqdec => //=. ce_solv.
   - move => p0 p1 u0 u1 neu0 neu1 h ih.
     rewrite check_equal_proj_proj.
     move /negP /nandP. case.
@@ -147,6 +145,14 @@ Proof.
     sauto lqb:on inv:CoqEq, CoqEq_Neu.
     sauto lqb:on inv:CoqEq, CoqEq_Neu.
     sauto lqb:on inv:CoqEq, CoqEq_Neu.
+  - rewrite /tm_conf.
+    move => a b n n0 i. simp ce_prop.
+    move => _. inversion 1; subst => //=.
+    + destruct b => //=.
+    + destruct a => //=.
+    + destruct b => //=.
+    + destruct a => //=.
+    + hauto l:on inv:CoqEq_Neu.
   - move => a b h ih.
     rewrite check_equal_nfnf.
     move : ih => /[apply].

From b29d567ef034039eaaae5880e086a25afd9f6936 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 5 Mar 2025 00:18:28 -0500
Subject: [PATCH 24/24] Add termination

---
 theories/termination.v | 14 ++++++++++++++
 1 file changed, 14 insertions(+)
 create mode 100644 theories/termination.v

diff --git a/theories/termination.v b/theories/termination.v
new file mode 100644
index 0000000..61f1e1d
--- /dev/null
+++ b/theories/termination.v
@@ -0,0 +1,14 @@
+Require Import common Autosubst2.core Autosubst2.unscoped Autosubst2.syntax algorithmic fp_red executable.
+From Hammer Require Import Tactics.
+Require Import ssreflect ssrbool.
+From stdpp Require Import relations (nsteps (..)).
+
+Definition term_metric k (a b : PTm) :=
+  exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j <= k.
+
+Lemma term_metric_algo_dom : forall k a b, term_metric k a b -> algo_dom_r a b.
+Proof.
+  elim /Wf_nat.lt_wf_ind.
+  move => n ih a b h.
+  case : (fancy_hred a); cycle 1.
+  move => [a' ha']. apply : A_HRedL; eauto.