From d925a8bcaa5c51d915600b3599c741690d0ad886 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 30 Jan 2025 20:23:57 -0500
Subject: [PATCH 001/210] Minor cleanup

---
 theories/fp_red.v | 10 ----------
 1 file changed, 10 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index a53a692..abc951e 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -144,16 +144,6 @@ with TRedSN {n} : PTm n -> PTm n -> Prop :=
 
 Derive Dependent Inversion tred_inv with (forall n (a b : PTm n), TRedSN a b) Sort Prop.
 
-Inductive SNe' {n} : PTm n -> Prop :=
-| N_Var' i :
-  SNe' (VarPTm i)
-| N_App' a b :
-  SNe a ->
-  SNe' (PApp a b)
-| N_Proj' p a :
-  SNe a ->
-  SNe' (PProj p a).
-
 Lemma PProjAbs_imp n p (a : PTm (S n)) :
   ~ SN (PProj p (PAbs a)).
 Proof.

From 9134cfec8a7c00177c0cd3cc4ca0c904b69688a6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 30 Jan 2025 22:18:58 -0500
Subject: [PATCH 002/210] Finish a few cases of eta postponement

---
 theories/fp_red.v | 140 ++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 140 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index abc951e..a255dd1 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -885,6 +885,12 @@ Module NeERed.
     move => h. elim : n a b /h=>//=; hauto lq:on ctrs:R_nonelim.
   Qed.
 
+  Lemma ToERed : forall n, (forall (a b : PTm n), R_elim a b -> ERed.R a b) /\
+                 (forall (a b : PTm n), R_nonelim a b -> ERed.R a b).
+  Proof.
+    apply ered_mutual; qauto l:on ctrs:ERed.R.
+  Qed.
+
 End NeERed.
 
 Module Type NoForbid.
@@ -1081,6 +1087,140 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto lq:on ctrs:rtc, NeERed.R_nonelim.
   Qed.
 
+
+  Lemma eta_postponement n a b c :
+    @P n a ->
+    ERed.R a b ->
+    RRed.R b c ->
+    exists d, rtc RRed.R a d /\ ERed.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.
+      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:ERed.R.
+    - move => n 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:ERed.R use:RReds.PairCong, RReds.ProjCong.
+    - move => n 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:ERed.R.
+    - move => n 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.
+        move {iha ihb}.
+        move /η_split /(_ hP') : ha.
+        move => [b [h0 h1]].
+        inversion h1; subst.
+        * inversion H0; subst.
+          exists (subst_PTm (scons b0 VarPTm) a3).
+          split; last by scongruence use:ERed.morphing.
+          apply : relations.rtc_transitive.
+          apply RReds.AppCong.
+          eassumption.
+          apply rtc_refl.
+          apply : rtc_l.
+          apply RRed.AppCong0. apply RRed.AbsCong. simpl. apply RRed.AppAbs.
+          asimpl.
+          apply rtc_once.
+          apply RRed.AppAbs.
+        * exfalso.
+          move : hP h0. clear => hP h0.
+          have : rtc RRed.R (PApp a0 b0) (PApp (PPair (PProj PL a1) (PProj PR a1)) b0)
+            by qauto l:on ctrs:rtc use:RReds.AppCong.
+          move : P_RReds hP. repeat move/[apply].
+          sfirstorder use:P_AppPair.
+        * exists (subst_PTm (scons b0 VarPTm) a1).
+          split.
+          apply : rtc_r; last by apply RRed.AppAbs.
+          hauto lq:on ctrs:rtc use:RReds.AppCong.
+          hauto l:on inv:option use:ERed.morphing,NeERed.ToERed.
+      + move => a2 a3 b2 ha2 [*]. subst.
+        move : iha (ha2) {ihb} => /[apply].
+        have : P a0 by sfirstorder use:P_AppInv.
+        move /[swap]/[apply].
+        move => [d [h0 h1]].
+        exists (PApp d b0).
+        hauto lq:on ctrs:ERed.R, rtc use:RReds.AppCong.
+      + move => a2 b0 b1 hb [*]. subst.
+        sauto lq:on.
+    - move => n a b0 b1 hb ihb Γ c A hu hu'.
+      elim /RRed.inv : hu' => //=_.
+      + move => A0 a0 b2 [*]. subst.
+        admit.
+      + sauto lq:on.
+      + move => a0 b2 b3 hb0 [*]. subst.
+        have [? [? ]] : exists Γ  A, @Wt n Γ b0 A by hauto lq:on inv:Wt.
+        move : ihb hb0. repeat move/[apply].
+        move => [d [h0 h1]].
+        exists (PApp a d).
+        split. admit.
+        sauto lq:on.
+    - move => n a0 a1 b ha iha Γ u A hu.
+      elim / RRed.inv => //= _.
+      + move => a2 a3 b0 h [*]. subst.
+        have [? [? ]] : exists Γ  A, @Wt n Γ a0 A by hauto lq:on inv:Wt.
+        move : iha h. repeat move/[apply].
+        move => [d [h0 h1]].
+        exists (PPair d b).
+        split. admit.
+        sauto lq:on.
+      + move => a2 b0 b1 h [*]. subst.
+        sauto lq:on.
+    - move => n a b0 b1 hb ihb Γ c A hu.
+      elim / RRed.inv => //=_.
+      move => a0 a1 b2 ha [*]. subst.
+      + sauto lq:on.
+      + move => a0 b2 b3 hb0 [*]. subst.
+        have [? [? ]] : exists Γ  A, @Wt n Γ b0 A by hauto lq:on inv:Wt.
+        move : ihb hb0. repeat move/[apply].
+        move => [d [h0 h1]].
+        exists (PPair a d).
+        split. admit.
+        sauto lq:on.
+    - move => n p a0 a1 ha iha Γ u A hu.
+      elim / RRed.inv => //=_.
+      + move => p0 a2 b0 [*]. subst.
+        inversion ha; subst.
+        * exfalso.
+          move : hu. clear. hauto q:on inv:Wt.
+        * exists (match p with
+             | PL => a2
+             | PR => b0
+             end).
+          split. apply : rtc_l.
+          apply RRed.ProjPair.
+          apply rtc_once. clear.
+          hauto lq:on use:RRed.ProjPair.
+          admit.
+        * eexists.
+          split. apply rtc_once.
+          apply RRed.ProjPair.
+          admit.
+        * eexists.
+          split. apply rtc_once.
+          apply RRed.ProjPair.
+          admit.
+      + move => p0 a2 a3 ha0 [*]. subst.
+        have [? [? ]] : exists Γ  A, @Wt n Γ a0 A by hauto lq:on inv:Wt.
+        move : iha ha0; repeat move/[apply].
+        move => [d [h0 h1]].
+        exists (PProj p d).
+        split.
+        admit.
+        sauto lq:on.
+  Admitted.
+
+
 End UniqueNF.
 
 Lemma η_nf_to_ne n (a0 a1 : PTm n) :

From 51ac5ffbd666c24c4eb16fded977a2e064bb431d Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Thu, 30 Jan 2025 23:10:11 -0500
Subject: [PATCH 003/210] Finish eta postponement

---
 theories/fp_red.v | 99 +++++++++++++++++------------------------------
 1 file changed, 35 insertions(+), 64 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index a255dd1..bc8ec58 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1151,75 +1151,46 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         move => [d [h0 h1]].
         exists (PApp d b0).
         hauto lq:on ctrs:ERed.R, rtc use:RReds.AppCong.
-      + move => a2 b0 b1 hb [*]. subst.
-        sauto lq:on.
-    - move => n a b0 b1 hb ihb Γ c A hu hu'.
-      elim /RRed.inv : hu' => //=_.
-      + move => A0 a0 b2 [*]. subst.
-        admit.
-      + sauto lq:on.
-      + move => a0 b2 b3 hb0 [*]. subst.
-        have [? [? ]] : exists Γ  A, @Wt n Γ b0 A by hauto lq:on inv:Wt.
-        move : ihb hb0. repeat move/[apply].
-        move => [d [h0 h1]].
-        exists (PApp a d).
-        split. admit.
-        sauto lq:on.
-    - move => n a0 a1 b ha iha Γ u A hu.
+      + move => a2 b2 b3 hb2 [*]. subst.
+        move {iha}.
+        have : P b0 by sfirstorder use:P_AppInv.
+        move : ihb hb2; repeat move /[apply].
+        hauto lq:on rew:off ctrs:ERed.R, rtc use:RReds.AppCong.
+    - move => n a0 a1 b0 b1 ha iha hb ihb c /P_PairInv [hP hP'].
+      elim /RRed.inv => //=_;
+        hauto lq:on rew:off ctrs:ERed.R, rtc use:RReds.PairCong.
+    - move => n p a0 a1 ha iha c /[dup] hP /P_ProjInv hP'.
       elim / RRed.inv => //= _.
-      + move => a2 a3 b0 h [*]. subst.
-        have [? [? ]] : exists Γ  A, @Wt n Γ a0 A by hauto lq:on inv:Wt.
-        move : iha h. repeat move/[apply].
-        move => [d [h0 h1]].
-        exists (PPair d b).
-        split. admit.
-        sauto lq:on.
-      + move => a2 b0 b1 h [*]. subst.
-        sauto lq:on.
-    - move => n a b0 b1 hb ihb Γ c A hu.
-      elim / RRed.inv => //=_.
-      move => a0 a1 b2 ha [*]. subst.
-      + sauto lq:on.
-      + move => a0 b2 b3 hb0 [*]. subst.
-        have [? [? ]] : exists Γ  A, @Wt n Γ b0 A by hauto lq:on inv:Wt.
-        move : ihb hb0. repeat move/[apply].
-        move => [d [h0 h1]].
-        exists (PPair a d).
-        split. admit.
-        sauto lq:on.
-    - move => n p a0 a1 ha iha Γ u A hu.
-      elim / RRed.inv => //=_.
       + move => p0 a2 b0 [*]. subst.
-        inversion ha; subst.
-        * exfalso.
-          move : hu. clear. hauto q:on inv:Wt.
-        * exists (match p with
-             | PL => a2
-             | PR => b0
-             end).
-          split. apply : rtc_l.
+        move : η_split hP'  ha; repeat move/[apply].
+        move => [a1 [h0 h1]].
+        inversion h1; subst.
+        * qauto l:on ctrs:rtc use:RReds.ProjCong, P_ProjAbs, P_RReds.
+        * inversion H0; subst.
+          exists (if p is PL then a1 else b1).
+          split; last by scongruence use:NeERed.ToERed.
+          apply : relations.rtc_transitive.
+          apply RReds.ProjCong; eauto.
+          apply : rtc_l.
+          apply RRed.ProjCong.
+          apply RRed.PairCong0.
           apply RRed.ProjPair.
-          apply rtc_once. clear.
-          hauto lq:on use:RRed.ProjPair.
-          admit.
-        * eexists.
-          split. apply rtc_once.
+          apply : rtc_l.
+          apply RRed.ProjCong.
+          apply RRed.PairCong1.
           apply RRed.ProjPair.
-          admit.
-        * eexists.
-          split. apply rtc_once.
+          apply rtc_once. apply RRed.ProjPair.
+        * exists (if p is PL then a3 else b1).
+          split; last by hauto lq:on use:NeERed.ToERed.
+          apply : relations.rtc_transitive.
+          eauto using RReds.ProjCong.
+          apply rtc_once.
           apply RRed.ProjPair.
-          admit.
-      + move => p0 a2 a3 ha0 [*]. subst.
-        have [? [? ]] : exists Γ  A, @Wt n Γ a0 A by hauto lq:on inv:Wt.
-        move : iha ha0; repeat move/[apply].
-        move => [d [h0 h1]].
-        exists (PProj p d).
-        split.
-        admit.
-        sauto lq:on.
-  Admitted.
-
+      + move => p0 a2 a3 h0 [*]. subst.
+        move : iha hP' h0;repeat move/[apply].
+        hauto lq:on ctrs:rtc, ERed.R use:RReds.ProjCong.
+    - hauto lq:on inv:RRed.R.
+  Qed.
 
 End UniqueNF.
 

From 5a7f46a8a19e48b5ac3b322c38c49418348eefa2 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Thu, 30 Jan 2025 23:29:25 -0500
Subject: [PATCH 004/210] Finish the enhanced eta postponement

---
 theories/fp_red.v | 30 ++++++++++++++++++++++++++++++
 1 file changed, 30 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index bc8ec58..20929e1 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1192,6 +1192,36 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto lq:on inv:RRed.R.
   Qed.
 
+  Lemma η_postponement_star n a b c :
+    @P n a ->
+    ERed.R a b ->
+    rtc RRed.R b c ->
+    exists d, rtc RRed.R a d /\ ERed.R d c.
+  Proof.
+    move => + + h. move : a.
+    elim : b c / h.
+    - sfirstorder.
+    - move => a0 a1 a2 ha ha' iha u hu hu'.
+      move : eta_postponement (hu) ha hu'; repeat move/[apply].
+      move => [d [h0 h1]].
+      have : P d by sfirstorder use:P_RReds.
+      move : iha h1; repeat move/[apply].
+      sfirstorder use:@relations.rtc_transitive.
+  Qed.
+
+  Lemma η_postponement_star' n a b c :
+    @P n a ->
+    ERed.R a b ->
+    rtc RRed.R b c ->
+    exists d, rtc RRed.R a d /\ NeERed.R_nonelim d c.
+  Proof.
+    move => h0 h1 h2.
+    have : exists d, rtc RRed.R a d /\ ERed.R d c by eauto using η_postponement_star.
+    move => [d [h3 /η_split]].
+    move /(_ ltac:(eauto using P_RReds)).
+    sfirstorder use:@relations.rtc_transitive.
+  Qed.
+
 End UniqueNF.
 
 Lemma η_nf_to_ne n (a0 a1 : PTm n) :

From 571779a4dd323ff254197ae2e456b004b3721ec0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 31 Jan 2025 17:32:35 -0500
Subject: [PATCH 005/210] Fix name change errors for EPar

---
 theories/fp_red.v | 38 +++++++++++++++++++-------------------
 1 file changed, 19 insertions(+), 19 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index a7624e2..2317fe2 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -888,7 +888,7 @@ Module NeEPar.
   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).
   Proof.
-    apply ered_mutual; qauto l:on ctrs:EPar.R.
+    apply epar_mutual; qauto l:on ctrs:EPar.R.
   Qed.
 
 End NeEPar.
@@ -1092,9 +1092,9 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
 
   Lemma eta_postponement n a b c :
     @P n a ->
-    ERed.R a b ->
+    EPar.R a b ->
     RRed.R b c ->
-    exists d, rtc RRed.R a d /\ ERed.R d c.
+    exists d, rtc RRed.R a d /\ EPar.R d c.
   Proof.
     move => + h.
     move : c.
@@ -1104,17 +1104,17 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       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:ERed.R.
+      hauto lq:on ctrs:EPar.R.
     - move => n 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:ERed.R use:RReds.PairCong, RReds.ProjCong.
+      hauto lq:on ctrs:EPar.R use:RReds.PairCong, RReds.ProjCong.
     - move => n 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:ERed.R.
+      hauto lq:on use:RReds.AbsCong ctrs:EPar.R.
     - move => n a0 a1 b0 b1 ha iha hb ihb c hP.
       elim /RRed.inv => //= _.
       + move => a2 b2 [*]. subst.
@@ -1125,7 +1125,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         inversion h1; subst.
         * inversion H0; subst.
           exists (subst_PTm (scons b0 VarPTm) a3).
-          split; last by scongruence use:ERed.morphing.
+          split; last by scongruence use:EPar.morphing.
           apply : relations.rtc_transitive.
           apply RReds.AppCong.
           eassumption.
@@ -1145,22 +1145,22 @@ 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:ERed.morphing,NeERed.ToERed.
+          hauto l:on inv:option use:EPar.morphing,NeEPar.ToEPar.
       + move => a2 a3 b2 ha2 [*]. subst.
         move : iha (ha2) {ihb} => /[apply].
         have : P a0 by sfirstorder use:P_AppInv.
         move /[swap]/[apply].
         move => [d [h0 h1]].
         exists (PApp d b0).
-        hauto lq:on ctrs:ERed.R, rtc use:RReds.AppCong.
+        hauto lq:on ctrs:EPar.R, rtc use:RReds.AppCong.
       + move => a2 b2 b3 hb2 [*]. subst.
         move {iha}.
         have : P b0 by sfirstorder use:P_AppInv.
         move : ihb hb2; repeat move /[apply].
-        hauto lq:on rew:off ctrs:ERed.R, rtc use:RReds.AppCong.
+        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'].
       elim /RRed.inv => //=_;
-        hauto lq:on rew:off ctrs:ERed.R, rtc use:RReds.PairCong.
+        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'.
       elim / RRed.inv => //= _.
       + move => p0 a2 b0 [*]. subst.
@@ -1170,7 +1170,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         * qauto l:on ctrs:rtc use:RReds.ProjCong, P_ProjAbs, P_RReds.
         * inversion H0; subst.
           exists (if p is PL then a1 else b1).
-          split; last by scongruence use:NeERed.ToERed.
+          split; last by scongruence use:NeEPar.ToEPar.
           apply : relations.rtc_transitive.
           apply RReds.ProjCong; eauto.
           apply : rtc_l.
@@ -1183,22 +1183,22 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           apply RRed.ProjPair.
           apply rtc_once. apply RRed.ProjPair.
         * exists (if p is PL then a3 else b1).
-          split; last by hauto lq:on use:NeERed.ToERed.
+          split; last by hauto lq:on use:NeEPar.ToEPar.
           apply : relations.rtc_transitive.
           eauto using RReds.ProjCong.
           apply rtc_once.
           apply RRed.ProjPair.
       + move => p0 a2 a3 h0 [*]. subst.
         move : iha hP' h0;repeat move/[apply].
-        hauto lq:on ctrs:rtc, ERed.R use:RReds.ProjCong.
+        hauto lq:on ctrs:rtc, EPar.R use:RReds.ProjCong.
     - hauto lq:on inv:RRed.R.
   Qed.
 
   Lemma η_postponement_star n a b c :
     @P n a ->
-    ERed.R a b ->
+    EPar.R a b ->
     rtc RRed.R b c ->
-    exists d, rtc RRed.R a d /\ ERed.R d c.
+    exists d, rtc RRed.R a d /\ EPar.R d c.
   Proof.
     move => + + h. move : a.
     elim : b c / h.
@@ -1213,12 +1213,12 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
 
   Lemma η_postponement_star' n a b c :
     @P n a ->
-    ERed.R a b ->
+    EPar.R a b ->
     rtc RRed.R b c ->
-    exists d, rtc RRed.R a d /\ NeERed.R_nonelim d c.
+    exists d, rtc RRed.R a d /\ NeEPar.R_nonelim d c.
   Proof.
     move => h0 h1 h2.
-    have : exists d, rtc RRed.R a d /\ ERed.R d c by eauto using η_postponement_star.
+    have : exists d, rtc RRed.R a d /\ EPar.R d c by eauto using η_postponement_star.
     move => [d [h3 /η_split]].
     move /(_ ltac:(eauto using P_RReds)).
     sfirstorder use:@relations.rtc_transitive.

From f57e10be93bed98aa1097ad67e0a5a41bcc5944b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 31 Jan 2025 20:10:56 -0500
Subject: [PATCH 006/210] Finish extracting leftmost-outermost red from sn

---
 theories/fp_red.v | 122 ++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 122 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 2317fe2..0ca4dfe 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -9,6 +9,8 @@ Require Import Psatz.
 From stdpp Require Import relations (rtc (..), rtc_once, rtc_r).
 From Hammer Require Import Tactics.
 Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+Require Import Btauto.
+Require Import Cdcl.Itauto.
 
 Ltac2 spec_refl () :=
   List.iter
@@ -1317,3 +1319,123 @@ Module RERed.
   Qed.
 
 End RERed.
+
+Module LoRed.
+  Inductive R {n} : PTm n -> PTm n ->  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)
+
+  (*************** Congruence ********************)
+  | AbsCong a0 a1 :
+    R a0 a1 ->
+    R (PAbs a0) (PAbs a1)
+  | AppCong0 a0 a1 b :
+    ~~ ishf a0 ->
+    R a0 a1 ->
+    R (PApp a0 b) (PApp a1 b)
+  | AppCong1 a b0 b1 :
+    ne a ->
+    R b0 b1 ->
+    R (PApp a b0) (PApp a b1)
+  | PairCong0 a0 a1 b :
+    R a0 a1 ->
+    R (PPair a0 b) (PPair a1 b)
+  | PairCong1 a b0 b1 :
+    nf a ->
+    R b0 b1 ->
+    R (PPair a b0) (PPair a b1)
+  | ProjCong p a0 a1 :
+    ~~ ishf a0 ->
+    R a0 a1 ->
+    R (PProj p a0) (PProj p a1).
+
+  Lemma hf_preservation n (a b : PTm n) :
+    LoRed.R a b ->
+    ishf a ->
+    ishf b.
+  Proof.
+    move => h. elim : n a b /h=>//=.
+  Qed.
+
+End LoRed.
+
+Module LoReds.
+  Lemma hf_preservation n (a b : PTm n) :
+    rtc LoRed.R a b ->
+    ishf a ->
+    ishf b.
+  Proof.
+    induction 1; eauto using LoRed.hf_preservation.
+  Qed.
+
+  Lemma hf_ne_imp n (a b : PTm n) :
+    rtc LoRed.R a b ->
+    ne b ->
+    ~~ ishf a.
+  Proof.
+    move : hf_preservation. repeat move/[apply].
+    case : a; case : b => //=; itauto.
+  Qed.
+
+  #[local]Ltac solve_s_rec :=
+  move => *; eapply rtc_l; eauto;
+         hauto lq:on ctrs:LoRed.R, rtc use:hf_ne_imp.
+
+  #[local]Ltac solve_s :=
+    repeat (induction 1; last by solve_s_rec); (move => *; apply rtc_refl).
+
+  Lemma AbsCong n (a b : PTm (S n)) :
+    rtc LoRed.R a b ->
+    rtc LoRed.R (PAbs a) (PAbs b).
+  Proof. solve_s. Qed.
+
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+    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) :
+    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) :
+    rtc LoRed.R a0 a1 ->
+    ne a1 ->
+    rtc LoRed.R (PProj p a0) (PProj p a1).
+  Proof. solve_s. Qed.
+
+  Local Ltac triv := simpl in *; itauto.
+
+  Lemma FromSN : 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).
+  Proof.
+    apply sn_mutual.
+    - hauto lq:on ctrs:rtc.
+    - hauto lq:on rew:off use:LoReds.AppCong solve+:triv.
+    - hauto l:on use:LoReds.ProjCong solve+:triv.
+    - hauto q:on use:LoReds.PairCong solve+:triv.
+    - hauto q:on use:LoReds.AbsCong solve+:triv.
+    - sfirstorder use:ne_nf.
+    - hauto lq:on ctrs:rtc.
+    - qauto ctrs:LoRed.R.
+    - move => n 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.
+      have : ~~ ishf a by inversion h.
+      hauto lq:on ctrs:LoRed.R.
+  Qed.
+End LoReds.

From 580e3a8251d69288623adfe8368987f246c49f24 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 31 Jan 2025 20:54:33 -0500
Subject: [PATCH 007/210] Prove that ered is strongly normalizing

---
 theories/fp_red.v | 35 +++++++++++++++++++++++++++++++++--
 1 file changed, 33 insertions(+), 2 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 0ca4dfe..aa9b111 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -4,9 +4,9 @@ Import Ltac2.Notations.
 Import Ltac2.Control.
 Require Import ssreflect ssrbool.
 Require Import FunInd.
-Require Import Arith.Wf_nat.
+Require Import Arith.Wf_nat (well_founded_lt_compat).
 Require Import Psatz.
-From stdpp Require Import relations (rtc (..), rtc_once, rtc_r).
+From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn).
 From Hammer Require Import Tactics.
 Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
 Require Import Btauto.
@@ -1438,4 +1438,35 @@ Module LoReds.
       have : ~~ ishf a by inversion h.
       hauto lq:on ctrs:LoRed.R.
   Qed.
+
 End LoReds.
+
+
+Fixpoint size_PTm {n} (a : PTm n) :=
+  match a with
+  | VarPTm _ => 1
+  | PAbs a => 1 + 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 => 1 + Nat.add (size_PTm a) (size_PTm b)
+  end.
+
+Lemma size_PTm_ren n m (ξ : fin n -> fin m) a : size_PTm (ren_PTm ξ a) = size_PTm a.
+Proof.
+  move : m ξ. elim : n / a => //=; scongruence.
+Qed.
+
+#[export]Hint Rewrite size_PTm_ren : sizetm.
+
+Lemma ered_size {n} (a b : PTm n) :
+  ERed.R a b ->
+  size_PTm b < size_PTm a.
+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.
+Proof.
+  hauto lq:on rew:off use:size_PTm_ren, ered_size,
+          well_founded_lt_compat unfold:well_founded.
+Qed.

From d2cd3105c7c1bac6d06147656c134b0353ecef4d Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 31 Jan 2025 21:45:55 -0500
Subject: [PATCH 008/210] stuck on antirenaming because of scoped syntax

---
 theories/fp_red.v | 128 ++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 128 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index aa9b111..13e66f8 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1260,12 +1260,52 @@ Module ERed.
     R a0 a1 ->
     R (PProj p a0) (PProj p a1).
 
+  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
+
   Lemma ToEPar n (a b : PTm n) :
     ERed.R a b -> EPar.R a b.
   Proof.
     induction 1; hauto lq:on use:EPar.refl ctrs:EPar.R.
   Qed.
 
+  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 ?x -> _ ?y => (ltac1:(case;qauto 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 antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
+    R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
+  Proof.
+    move E : (ren_PTm ξ a) => u hu.
+    move : n ξ a E.
+    elim : m u b / hu; try solve_anti_ren.
+    - move => n a m ξ []//=.
+      move => u [].
+      case : u => //=.
+      move => u0 u1 [].
+      case : u1 => //=.
+      move => i /[swap] [].
+      case : i => //= _ h.
+      apply f_equal with (f := subst_PTm (scons (PAbs (VarPTm var_zero)) VarPTm)) in h.
+      move : h. asimpl. move => ?. subst.
+
+      rewrite -/var_zero.
+      eexists. split. apply AppEta.
+      move : h. asimpl => ?.  subst.
+
+
+
+
 End ERed.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -1470,3 +1510,91 @@ 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) :
+  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.
+  - move => a c.
+    elim /ERed.inv => //= _.
+    + move => a0 [+ ?]. subst => h.
+      apply f_equal with (f := subst_PTm (scons (PAbs (VarPTm var_zero)) VarPTm)) in h.
+      move : h. asimpl => ?. subst.
+      eauto using rtc_refl.
+    + move => a0 a1 ha [*]. subst.
+      elim /ERed.inv : ha => //= _.
+      * move => a0 a2 b0 ha [*]. subst. rename a2 into a1.
+        have [a2 [h0 h1]] : exists a2, ERed.R a2 a /\ a1 = ren_PTm shift a2 by admit. subst.
+        eexists. split; cycle 1.
+        apply : relations.rtc_r; cycle 1.
+        apply ERed.AppEta.
+        apply rtc_refl.
+        eauto using relations.rtc_once.
+      * hauto q:on ctrs:rtc, ERed.R inv:ERed.R.
+  - move => a c ha.
+    elim /ERed.inv : ha => //= _.
+    + hauto l:on.
+    + move => a0 a1 b0 ha ? [*]. subst.
+      elim /ERed.inv : ha => //= _.
+      move => p a1 a2 ha ? [*]. subst.
+      exists a1. split. by apply relations.rtc_once.
+      apply : rtc_l. apply ERed.PairEta.
+      apply : rtc_l. apply ERed.PairCong1. eauto using ERed.ProjCong.
+      apply rtc_refl.
+    + move => a0 b0 b1 ha ? [*]. subst.
+      elim /ERed.inv : ha => //= _ p a0 a1 h ? [*]. subst.
+      exists a0. split; first by apply relations.rtc_once.
+      apply : rtc_l; first by apply ERed.PairEta.
+      apply relations.rtc_once.
+      hauto lq:on ctrs:ERed.R.
+  - move => a0 a1 ha iha c.
+    elim /ERed.inv => //= _.
+    + move => a2 ? [*]. subst.
+      elim /ERed.inv : ha => //=_.
+      * move => a1 a2 b0 ha ? [*] {iha}. subst.
+        have [a0 [h0 h1]] : exists a0, ERed.R a0 c /\ a1 = ren_Tm shift a0 by admit. subst.
+        exists a0. split; last by apply relations.rtc_once.
+        apply relations.rtc_once. apply ERed.AppEta.
+      * hauto q:on inv:ERed.R.
+    + hauto l:on use:EReds.AbsCong.
+  - move => a0 a1 b ha iha c.
+    elim /ERed.inv => //= _.
+    + hauto lq:on ctrs:rtc use:EReds.AppCong.
+    + hauto lq:on use:@relations.rtc_once ctrs:ERed.R.
+  - move => a b0 b1 hb ihb c.
+    elim /ERed.inv => //=_.
+    + move => a0 a1 a2 ha ? [*]. subst.
+      move {ihb}. exists (App a0 b0).
+      hauto lq:on use:@relations.rtc_once ctrs:ERed.R.
+    + hauto lq:on ctrs:rtc use:EReds.AppCong.
+  - move => a0 a1 b ha iha c.
+    elim /ERed.inv => //= _.
+    + move => ? ?[*]. subst.
+      elim /ERed.inv : ha => //= _ p a1 a2 h ? [*]. subst.
+      exists a1. split; last by apply relations.rtc_once.
+      apply : rtc_l. apply ERed.PairEta.
+      apply relations.rtc_once. hauto lq:on ctrs:ERed.R.
+    + hauto lq:on ctrs:rtc use:EReds.PairCong.
+    + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+  - move => a b0 b1 hb hc c. elim /ERed.inv => //= _.
+    + move => ? ? [*]. subst.
+      elim /ERed.inv : hb => //= _ p a0 a1 ha ? [*]. subst.
+      move {hc}.
+      exists a0. split; last by apply relations.rtc_once.
+      apply : rtc_l; first by apply ERed.PairEta.
+      hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+    + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+    + hauto lq:on ctrs:rtc use:EReds.PairCong.
+  - qauto l:on inv:ERed.R use:EReds.ProjCong.
+  - move => p A0 A1 B hA ihA.
+    move => c. elim/ERed.inv => //=.
+    + hauto lq:on ctrs:rtc use:EReds.BindCong.
+    + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+  - move => p A B0 B1 hB ihB c.
+    elim /ERed.inv => //=.
+    + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+    + hauto lq:on ctrs:rtc use:EReds.BindCong.
+Admitted.

From ecb50f1ab7e78905a6afb3a46a56e7bab9a6b7f3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 31 Jan 2025 21:56:58 -0500
Subject: [PATCH 009/210] Add comment about the antirenaming proof

---
 theories/fp_red.v | 12 ++++++++----
 1 file changed, 8 insertions(+), 4 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 13e66f8..85a9512 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1282,7 +1282,7 @@ Module ERed.
   (*   destruct a. *)
   (*   exact (ξ f). *)
 
-
+  (* Need to generalize to injective renaming *)
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
     R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
   Proof.
@@ -1298,10 +1298,14 @@ Module ERed.
       case : i => //= _ h.
       apply f_equal with (f := subst_PTm (scons (PAbs (VarPTm var_zero)) VarPTm)) in h.
       move : h. asimpl. move => ?. subst.
+      admit.
+    - move => n a m ξ []//=.
+      move => u u0 [].
+      case : u => //=.
+      case : u0 => //=.
+      move => p p0 p1 p2 [? ?] [? ?]. subst.
+
 
-      rewrite -/var_zero.
-      eexists. split. apply AppEta.
-      move : h. asimpl => ?.  subst.
 
 
 

From 6cc3a651630f3e0698865bccc708db864e023a5e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 2 Feb 2025 19:42:42 -0500
Subject: [PATCH 010/210] Discharge one case of antirenaming using injective
 renaming

---
 theories/Autosubst2/syntax.v |  4 ++--
 theories/fp_red.v            | 36 ++++++++++++++++++++++++++++++------
 2 files changed, 32 insertions(+), 8 deletions(-)

diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index b5cbc66..78c4fec 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -804,9 +804,9 @@ Arguments PApp {n_PTm}.
 
 Arguments PAbs {n_PTm}.
 
-#[global]Hint Opaque subst_PTm: rewrite.
+#[global] Hint Opaque subst_PTm: rewrite.
 
-#[global]Hint Opaque ren_PTm: rewrite.
+#[global] Hint Opaque ren_PTm: rewrite.
 
 End Extra.
 
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 85a9512..428669f 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1282,28 +1282,52 @@ Module ERed.
   (*   destruct a. *)
   (*   exact (ξ f). *)
 
+  Lemma ren_injective n m (a b : PTm n) (ξ : fin n -> fin m) :
+    (forall i j, ξ i = ξ j -> i = j) ->
+    ren_PTm ξ a = ren_PTm ξ b ->
+    a = b.
+  Proof.
+    move : m ξ b.
+    elim : n / a => //; try solve_anti_ren.
+
+    move => n a iha m ξ []//=.
+    move => u hξ [h].
+    apply iha in h. by subst.
+    destruct i, j=>//=.
+    hauto l:on.
+  Qed.
+
+(* forall i j : fin m, ξ i = ξ j -> i = j *)
+
   (* Need to generalize to injective renaming *)
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : 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.
   Proof.
+    move => hξ.
     move E : (ren_PTm ξ a) => u hu.
-    move : n ξ a E.
+    move : n ξ hξ a E.
     elim : m u b / hu; try solve_anti_ren.
-    - move => n a m ξ []//=.
+    - move => n a m ξ hξ []//=.
       move => u [].
       case : u => //=.
       move => u0 u1 [].
       case : u1 => //=.
       move => i /[swap] [].
       case : i => //= _ h.
-      apply f_equal with (f := subst_PTm (scons (PAbs (VarPTm var_zero)) VarPTm)) in h.
-      move : h. asimpl. move => ?. subst.
+      move : h. asimpl.
+      apply f_equal with (f := ren_PTm (scons var_zero id)) in h.
+      move : h. asimpl.
+      Check scons var_zero shift.
       admit.
-    - move => n a m ξ []//=.
+    - move => n a m ξ hξ []//=.
       move => u u0 [].
       case : u => //=.
       case : u0 => //=.
-      move => p p0 p1 p2 [? ?] [? ?]. subst.
+      move => p p0 p1 p2 [? ?] [? h]. subst.
+      have ? : p0 = p2 by eauto using ren_injective. subst.
+      hauto l:on.
+    -
 
 
 

From 5624415bc96e766911c479e878645f958cfd6bd9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 2 Feb 2025 20:08:37 -0500
Subject: [PATCH 011/210] Finish antirenaming for injective rens

---
 theories/fp_red.v | 39 ++++++++++++++++++++++++---------------
 1 file changed, 24 insertions(+), 15 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 428669f..6e09b98 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1272,7 +1272,7 @@ 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 ?x -> _ ?y => (ltac1:(case;qauto depth:2 ctrs:ERed.R))
+    | fin ?x -> _ ?y => (ltac1:(case;hauto q:on depth:2 ctrs:ERed.R))
     | _ => solve_anti_ren ()
     end.
 
@@ -1282,6 +1282,13 @@ Module ERed.
   (*   destruct a. *)
   (*   exact (ξ f). *)
 
+  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.
+  Proof.
+    sblast inv:option.
+  Qed.
+
   Lemma ren_injective n m (a b : PTm n) (ξ : fin n -> fin m) :
     (forall i j, ξ i = ξ j -> i = j) ->
     ren_PTm ξ a = ren_PTm ξ b ->
@@ -1306,32 +1313,34 @@ Module ERed.
   Proof.
     move => hξ.
     move E : (ren_PTm ξ a) => u hu.
-    move : n ξ hξ a E.
+    move : n ξ a hξ E.
     elim : m u b / hu; try solve_anti_ren.
-    - move => n a m ξ hξ []//=.
-      move => u [].
+    - move => n a m ξ []//=.
+      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.
+      move => [p ?]. subst.
       move : h. asimpl.
-      apply f_equal with (f := ren_PTm (scons var_zero id)) in h.
-      move : h. asimpl.
-      Check scons var_zero shift.
-      admit.
-    - move => n a m ξ hξ []//=.
-      move => u u0 [].
+      replace (ren_PTm (funcomp shift ξ) p) with
+        (ren_PTm shift (ren_PTm ξ p)); last by asimpl.
+      move /ren_injective.
+      move /(_ ltac:(hauto l:on)).
+      move => ?. subst.
+      exists p. split=>//. apply AppEta.
+    - move => n a m ξ [] //=.
+      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.
+  Admitted.
 
 
 End ERed.

From 376fce619cbc80e73ffccf2004609a3e94314512 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 2 Feb 2025 20:21:06 -0500
Subject: [PATCH 012/210] Save progress

---
 theories/fp_red.v | 75 +++++++++++++++++++++++++++++++++--------------
 1 file changed, 53 insertions(+), 22 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 6e09b98..413c34a 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -773,6 +773,45 @@ Module RReds.
 
 End RReds.
 
+Module EReds.
+
+  #[local]Ltac solve_s_rec :=
+  move => *; eapply rtc_l; eauto;
+         hauto lq:on ctrs:ERed.R.
+
+  #[local]Ltac solve_s :=
+    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
+
+  Lemma AbsCong n (a b : PTm (S n)) :
+    rtc ERed.R a b ->
+    rtc ERed.R (PAbs a) (PAbs b).
+  Proof. solve_s. Qed.
+
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+    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) :
+    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) :
+    rtc ERed.R a0 a1 ->
+    rtc ERed.R (PProj p a0) (PProj p a1).
+  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).
+  Proof.
+    move => h. move : m ξ. elim : a b /h; hauto lq:on ctrs:rtc use:RRed.renaming.
+  Qed.
+
+End RReds.
+
 
 Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
   (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)).
@@ -1564,39 +1603,31 @@ Proof.
     + move => a0 a1 ha [*]. subst.
       elim /ERed.inv : ha => //= _.
       * move => a0 a2 b0 ha [*]. subst. rename a2 into a1.
-        have [a2 [h0 h1]] : exists a2, ERed.R a2 a /\ a1 = ren_PTm shift a2 by admit. subst.
-        eexists. split; cycle 1.
-        apply : relations.rtc_r; cycle 1.
-        apply ERed.AppEta.
-        apply rtc_refl.
-        eauto using relations.rtc_once.
+        move /ERed.antirenaming : ha.
+        move /(_ ltac:(hauto lq:on)) => [a' [h0 h1]]. subst.
+        hauto lq:on ctrs:rtc, ERed.R.
       * hauto q:on ctrs:rtc, ERed.R inv:ERed.R.
   - move => a c ha.
     elim /ERed.inv : ha => //= _.
     + hauto l:on.
-    + move => a0 a1 b0 ha ? [*]. subst.
+    + move => a0 a1 b0 ha [*]. subst.
       elim /ERed.inv : ha => //= _.
-      move => p a1 a2 ha ? [*]. subst.
-      exists a1. split. by apply relations.rtc_once.
-      apply : rtc_l. apply ERed.PairEta.
-      apply : rtc_l. apply ERed.PairCong1. eauto using ERed.ProjCong.
-      apply rtc_refl.
-    + move => a0 b0 b1 ha ? [*]. subst.
-      elim /ERed.inv : ha => //= _ p a0 a1 h ? [*]. subst.
-      exists a0. split; first by apply relations.rtc_once.
-      apply : rtc_l; first by apply ERed.PairEta.
-      apply relations.rtc_once.
-      hauto lq:on ctrs:ERed.R.
+      move => p a0 a2 ha [*]. subst.
+      hauto q:on ctrs:rtc, ERed.R.
+    + move => a0 b0 b1 ha [*]. subst.
+      elim /ERed.inv : ha => //= _.
+      move => p a0 a2 ha [*]. subst.
+      hauto q:on ctrs:rtc, ERed.R.
   - move => a0 a1 ha iha c.
     elim /ERed.inv => //= _.
-    + move => a2 ? [*]. subst.
+    + move => a2 [*]. subst.
       elim /ERed.inv : ha => //=_.
-      * move => a1 a2 b0 ha ? [*] {iha}. subst.
-        have [a0 [h0 h1]] : exists a0, ERed.R a0 c /\ a1 = ren_Tm shift a0 by admit. subst.
+      * move => a0 a2 b0 ha [*] {iha}. subst.
+        have [a0 [h0 h1]] : exists a0, ERed.R c a0 /\ a2 = ren_PTm shift a0 by hauto lq:on use:ERed.antirenaming. subst.
         exists a0. split; last by apply relations.rtc_once.
         apply relations.rtc_once. apply ERed.AppEta.
       * hauto q:on inv:ERed.R.
-    + hauto l:on use:EReds.AbsCong.
+    + best use:EReds.AbsCong.
   - move => a0 a1 b ha iha c.
     elim /ERed.inv => //= _.
     + hauto lq:on ctrs:rtc use:EReds.AppCong.

From cf31e6d2854657a151ed4a63f2a928c097a38a43 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 2 Feb 2025 20:48:39 -0500
Subject: [PATCH 013/210] Finish the confluence proof of ered

---
 theories/fp_red.v | 141 ++++++++++++++++++++--------------------------
 1 file changed, 60 insertions(+), 81 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 413c34a..3e2982b 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -229,7 +229,7 @@ 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 ?x -> _ ?y => (ltac1:(case;qauto depth:2 db:sn))
+  | fin _ -> _ _ => (ltac1:(case;qauto depth:2 db:sn))
   | _ => solve_anti_ren ()
   end.
 
@@ -431,32 +431,25 @@ Module RRed.
     all : qauto ctrs:R.
   Qed.
 
+  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 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) :
     R (ren_PTm ξ a) b -> exists b0, R a b0 /\ ren_PTm ξ b0 = b.
   Proof.
     move E : (ren_PTm ξ a) => u h.
-    move : n ξ a E. elim : m u b/h.
+    move : n ξ a E. elim : m u b/h; try solve_anti_ren.
     - move => n a b m ξ []//=. move => []//= t t0 [*]. subst.
       eexists. split. apply AppAbs. by asimpl.
     - move => n p a b m ξ []//=.
       move => p0 []//=. hauto q:on ctrs:R.
-    - move => n a0 a1 ha iha m ξ []//=.
-      move => p [*]. subst.
-      spec_refl.
-      move : iha => [t [h0 h1]]. subst.
-      eexists. split. eauto using AbsCong.
-      by asimpl.
-    - move => n a0 a1 b ha iha m ξ []//=.
-      hauto lq:on ctrs:R.
-    - move => n a b0 b1 h ih m ξ []//=.
-      hauto lq:on ctrs:R.
-    - move => n a0 a1 b ha iha m ξ []//=.
-      hauto lq:on ctrs:R.
-    - move => n a b0 b1 h ih m ξ []//=.
-      hauto lq:on ctrs:R.
-    - move => n p a0 a1 ha iha m ξ []//=.
-      hauto lq:on ctrs:R.
   Qed.
 
 End RRed.
@@ -773,45 +766,6 @@ Module RReds.
 
 End RReds.
 
-Module EReds.
-
-  #[local]Ltac solve_s_rec :=
-  move => *; eapply rtc_l; eauto;
-         hauto lq:on ctrs:ERed.R.
-
-  #[local]Ltac solve_s :=
-    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
-
-  Lemma AbsCong n (a b : PTm (S n)) :
-    rtc ERed.R a b ->
-    rtc ERed.R (PAbs a) (PAbs b).
-  Proof. solve_s. Qed.
-
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
-    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) :
-    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) :
-    rtc ERed.R a0 a1 ->
-    rtc ERed.R (PProj p a0) (PProj p a1).
-  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).
-  Proof.
-    move => h. move : m ξ. elim : a b /h; hauto lq:on ctrs:rtc use:RRed.renaming.
-  Qed.
-
-End RReds.
-
 
 Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
   (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)).
@@ -1311,7 +1265,7 @@ 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 ?x -> _ ?y => (ltac1:(case;hauto q:on depth:2 ctrs:ERed.R))
+    | fin _ -> _ _ => (ltac1:(case;hauto q:on depth:2 ctrs:ERed.R))
     | _ => solve_anti_ren ()
     end.
 
@@ -1384,6 +1338,39 @@ Module ERed.
 
 End ERed.
 
+Module EReds.
+
+  #[local]Ltac solve_s_rec :=
+  move => *; eapply rtc_l; eauto;
+         hauto lq:on ctrs:ERed.R.
+
+  #[local]Ltac solve_s :=
+    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
+
+  Lemma AbsCong n (a b : PTm (S n)) :
+    rtc ERed.R a b ->
+    rtc ERed.R (PAbs a) (PAbs b).
+  Proof. solve_s. Qed.
+
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+    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) :
+    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) :
+    rtc ERed.R a0 a1 ->
+    rtc ERed.R (PProj p a0) (PProj p a1).
+  Proof. solve_s. Qed.
+
+End EReds.
+
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
 
 
@@ -1627,42 +1614,34 @@ Proof.
         exists a0. split; last by apply relations.rtc_once.
         apply relations.rtc_once. apply ERed.AppEta.
       * hauto q:on inv:ERed.R.
-    + best use:EReds.AbsCong.
+    + hauto lq:on use:EReds.AbsCong.
   - move => a0 a1 b ha iha c.
     elim /ERed.inv => //= _.
     + hauto lq:on ctrs:rtc use:EReds.AppCong.
     + hauto lq:on use:@relations.rtc_once ctrs:ERed.R.
   - move => a b0 b1 hb ihb c.
     elim /ERed.inv => //=_.
-    + move => a0 a1 a2 ha ? [*]. subst.
-      move {ihb}. exists (App a0 b0).
+    + move => a0 a1 a2 ha [*]. subst.
+      move {ihb}. exists (PApp a1 b1).
       hauto lq:on use:@relations.rtc_once ctrs:ERed.R.
     + hauto lq:on ctrs:rtc use:EReds.AppCong.
   - move => a0 a1 b ha iha c.
     elim /ERed.inv => //= _.
-    + move => ? ?[*]. subst.
-      elim /ERed.inv : ha => //= _ p a1 a2 h ? [*]. subst.
-      exists a1. split; last by apply relations.rtc_once.
-      apply : rtc_l. apply ERed.PairEta.
-      apply relations.rtc_once. hauto lq:on ctrs:ERed.R.
+    + sauto lq:on.
     + hauto lq:on ctrs:rtc use:EReds.PairCong.
     + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
   - move => a b0 b1 hb hc c. elim /ERed.inv => //= _.
-    + move => ? ? [*]. subst.
-      elim /ERed.inv : hb => //= _ p a0 a1 ha ? [*]. subst.
-      move {hc}.
-      exists a0. split; last by apply relations.rtc_once.
-      apply : rtc_l; first by apply ERed.PairEta.
-      hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+    + move => ? [*]. subst.
+      sauto lq:on.
     + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
     + hauto lq:on ctrs:rtc use:EReds.PairCong.
   - qauto l:on inv:ERed.R use:EReds.ProjCong.
-  - move => p A0 A1 B hA ihA.
-    move => c. elim/ERed.inv => //=.
-    + hauto lq:on ctrs:rtc use:EReds.BindCong.
-    + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
-  - move => p A B0 B1 hB ihB c.
-    elim /ERed.inv => //=.
-    + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
-    + hauto lq:on ctrs:rtc use:EReds.BindCong.
-Admitted.
+Qed.
+
+Lemma ered_confluence n (a b c : PTm n) :
+  rtc ERed.R a b ->
+  rtc ERed.R a c ->
+  exists d, rtc ERed.R b d  /\ rtc ERed.R c d.
+Proof.
+  sfirstorder use:relations.locally_confluent_confluent, ered_sn, ered_local_confluence.
+Qed.

From a0c20851fe2ca714b1f12305eaba5fcbec95f913 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 2 Feb 2025 21:57:41 -0500
Subject: [PATCH 014/210] Prove confluence of beta

---
 theories/fp_red.v | 119 ++++++++++++++++++++++++++++++----------------
 1 file changed, 79 insertions(+), 40 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3e2982b..0b6ef9c 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -567,6 +567,48 @@ Module RPar.
     induction 1; qauto l:on use:RPar.refl ctrs:RPar.R.
   Qed.
 
+  Function tstar {n} (a : PTm n) :=
+    match a with
+    | VarPTm i => a
+    | PAbs a => PAbs (tstar a)
+    | PApp (PAbs a) b => subst_PTm (scons (tstar b) VarPTm) (tstar a)
+    | PApp a b => PApp (tstar a) (tstar b)
+    | PPair a b => PPair (tstar a) (tstar b)
+    | PProj p (PPair a b) => if p is PL then (tstar a) else (tstar b)
+    | PProj p a => PProj p (tstar a)
+    end.
+
+  Lemma triangle n (a b : PTm n) :
+    RPar.R a b -> RPar.R b (tstar a).
+  Proof.
+    move : b.
+    apply tstar_ind => {}n{}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.
+    - hauto lq:on ctrs:R inv:R.
+    - hauto lq:on ctrs:R inv:R.
+    - move => p a0 b ? ? ih b0. subst.
+      elim /inv => //=_.
+      + move => p a1 a2 b1 b2 h0 h1[*]. subst.
+        by apply ih.
+      + move => p a1 a2 ha [*]. subst.
+        elim /inv : ha => //=_.
+        move => a1 a3 b0 b1 h0 h1 [*]. subst.
+        apply : ProjPair'; eauto using refl.
+    - move => p a0 b ? p0 ?. subst.
+      case : p0 => //= _.
+      move => ih b0. elim /inv => //=_.
+      + hauto l:on.
+      + move => p a1 a2 ha [*]. subst.
+        elim /inv : ha => //=_ > ? ? [*]. subst.
+        apply : ProjPair'; eauto using refl.
+    - hauto lq:on ctrs:R inv:R.
+  Qed.
+
+  Lemma diamond n (a b c : PTm n) :
+    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 :
@@ -603,46 +645,6 @@ Proof.
   - sauto.
 Qed.
 
-Function tstar {n} (a : PTm n) :=
-  match a with
-  | VarPTm i => a
-  | PAbs a => PAbs (tstar a)
-  | PApp (PAbs a) b => subst_PTm (scons (tstar b) VarPTm) (tstar a)
-  | PApp a b => PApp (tstar a) (tstar b)
-  | PPair a b => PPair (tstar a) (tstar b)
-  | PProj p (PPair a b) => if p is PL then (tstar a) else (tstar b)
-  | PProj p a => PProj p (tstar a)
-  end.
-
-Module TStar.
-  Lemma renaming n m (ξ : fin n -> fin m) (a : PTm n) :
-    tstar (ren_PTm ξ a) = ren_PTm ξ (tstar a).
-  Proof.
-    move : m ξ.
-    apply tstar_ind => {}n {}a => //=.
-    - hauto lq:on.
-    - scongruence.
-    - move => a0 b ? h ih m ξ.
-      rewrite ih.
-      asimpl; congruence.
-    - qauto l:on.
-    - scongruence.
-    - hauto q:on.
-    - qauto l:on.
-  Qed.
-
-  Lemma pair n (a b : PTm n) :
-    tstar (PPair a b) = PPair (tstar a) (tstar b).
-    reflexivity. Qed.
-End TStar.
-
-Definition isPair {n} (a : PTm n) := if a is PPair _ _ then true else false.
-
-Lemma tstar_proj n (a : PTm n) :
-  ((~~ isPair a) /\ forall p, tstar (PProj p a) = PProj p (tstar a)) \/
-    exists a0 b0, a = PPair a0 b0 /\ forall p, tstar (PProj p a) = (if p is PL then (tstar a0) else (tstar b0)).
-Proof. sauto lq:on. Qed.
-
 Module EPar'.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (****************** Eta ***********************)
@@ -764,6 +766,27 @@ Module RReds.
     move => h. move : m ξ. 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) :
+    rtc RRed.R a b.
+  Proof.
+    elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl.
+    move => n 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.
+    apply : rtc_r; last by apply RRed.ProjPair.
+    by eauto using PairCong, ProjCong.
+  Qed.
+
+  Lemma RParIff n (a b : PTm n) :
+    rtc RRed.R a b <-> rtc RPar.R a b.
+  Proof.
+    split.
+    induction 1; hauto l:on ctrs:rtc use:RPar.FromRRed, @relations.rtc_transitive.
+    induction 1; hauto l:on ctrs:rtc use:FromRPar, @relations.rtc_transitive.
+  Qed.
+
 End RReds.
 
 
@@ -1645,3 +1668,19 @@ Lemma ered_confluence n (a b c : PTm n) :
 Proof.
   sfirstorder use:relations.locally_confluent_confluent, ered_sn, ered_local_confluence.
 Qed.
+
+Lemma red_confluence n (a b c : PTm n) :
+  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
+             by hauto lq:on use:RReds.RParIff.
+  apply relations.diamond_confluent.
+  rewrite /relations.diamond.
+  eauto using RPar.diamond.
+Qed.
+
+(* "Declarative" Joinability *)
+Module DJoin.
+  Definition R {n} (a b : PTm n) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.
+End DJoin.

From f3f3991b02607a32daa15c49c9b6919535ea10cb Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Feb 2025 12:16:56 -0500
Subject: [PATCH 015/210] Finish rred standardization

---
 theories/fp_red.v | 32 ++++++++++++++++++++++++++++++++
 1 file changed, 32 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 0b6ef9c..926ab5d 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1444,6 +1444,12 @@ Module RERed.
     hauto q:on use:ERed.ToEPar, ToBetaEta.
   Qed.
 
+  Lemma sn_preservation n (a b : PTm n) :
+    R a b ->
+    SN a ->
+    SN b.
+  Proof. hauto q:on use:ToBetaEtaPar, epar_sn_preservation, red_sn_preservation, RPar.FromRRed. Qed.
+
 End RERed.
 
 Module LoRed.
@@ -1680,6 +1686,32 @@ Lemma red_confluence n (a b c : PTm n) :
   eauto using RPar.diamond.
 Qed.
 
+Lemma rered_standardization n (a c : PTm n) :
+  SN a ->
+  rtc RERed.R a c ->
+  exists b, rtc RRed.R a b /\ rtc NeEPar.R_nonelim b c.
+Proof.
+  move => + h. elim : a c /h.
+  by eauto using rtc_refl.
+  move => a b c.
+  move /RERed.ToBetaEtaPar.
+  case.
+  - move => h0 h1 ih hP.
+    have  : SN b by qauto use:epar_sn_preservation.
+    move => {}/ih [b' [ihb0 ihb1]].
+    hauto lq:on ctrs:rtc use:SN_UniqueNF.η_postponement_star'.
+  - hauto lq:on ctrs:rtc use:red_sn_preservation, RPar.FromRRed.
+Qed.
+
+Lemma rered_confluence n (a b c : PTm n) :
+  SN a ->
+  rtc RERed.R a b ->
+  rtc RERed.R a c ->
+  exists d, rtc RERed.R b d  /\ rtc RERed.R c d.
+Proof.
+  move => hP hb hc.
+
+
 (* "Declarative" Joinability *)
 Module DJoin.
   Definition R {n} (a b : PTm n) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.

From 3d42581bbe1591fc5a24e6921919cad62f254690 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Feb 2025 14:23:12 -0500
Subject: [PATCH 016/210] Finish the confluence proof

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

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 926ab5d..337cd11 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -452,6 +452,11 @@ Module RRed.
       move => p0 []//=. hauto q:on ctrs:R.
   Qed.
 
+  Lemma nf_imp n (a b : PTm n) :
+    nf a ->
+    R a b -> False.
+  Proof. move/[swap]. induction 1; hauto qb:on inv:PTm. Qed.
+
 End RRed.
 
 Module RPar.
@@ -727,6 +732,7 @@ Module EPars.
     rtc EPar'.R a0 a1 ->
     rtc EPar'.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
+
 End EPars.
 
 Module RReds.
@@ -787,6 +793,10 @@ Module RReds.
     induction 1; hauto l:on ctrs:rtc use:FromRPar, @relations.rtc_transitive.
   Qed.
 
+  Lemma nf_refl n (a b : PTm n) :
+    rtc RRed.R a b -> nf a -> a = b.
+  Proof. induction 1; sfirstorder use:RRed.nf_imp. Qed.
+
 End RReds.
 
 
@@ -1320,7 +1330,21 @@ Module ERed.
     hauto l:on.
   Qed.
 
-(* forall i j : fin m, ξ i = ξ j -> i = j *)
+  Lemma AppEta' n a u :
+    u = (@PApp (S n) (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) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
+  Proof.
+    move => h. move : m ξ.
+    elim : n a b /h.
+
+    move => n a m ξ /=.
+    apply AppEta'; eauto. by asimpl.
+    all : qauto ctrs:R.
+  Qed.
 
   (* Need to generalize to injective renaming *)
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
@@ -1392,6 +1416,24 @@ Module EReds.
     rtc ERed.R (PProj p a0) (PProj p a1).
   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).
+  Proof. induction 1; hauto l:on use:ERed.renaming ctrs:rtc. Qed.
+
+  Lemma FromEPar n (a b : PTm n) :
+    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.
+    - move => n 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.
+      apply : rtc_r.
+      apply PairCong; eauto using ProjCong.
+      apply ERed.PairEta.
+  Qed.
 End EReds.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -1437,6 +1479,14 @@ Module RERed.
     ERed.R a b \/ RRed.R a b.
   Proof. induction 1; hauto lq:on db:red. Qed.
 
+  Lemma FromBeta n (a b : PTm n) :
+    RRed.R a b -> RERed.R a b.
+  Proof. induction 1; qauto l:on ctrs:R. Qed.
+
+  Lemma FromEta n (a b : PTm n) :
+    ERed.R a b -> RERed.R a b.
+  Proof. induction 1; qauto l:on ctrs:R. Qed.
+
   Lemma ToBetaEtaPar n (a b : PTm n) :
     R a b ->
     EPar.R a b \/ RRed.R a b.
@@ -1452,6 +1502,25 @@ Module RERed.
 
 End RERed.
 
+Module REReds.
+  Lemma sn_preservation n (a b : PTm n) :
+    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) :
+    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) :
+    rtc ERed.R a b ->
+    rtc RERed.R a b.
+  Proof. induction 1; hauto lq:on ctrs:rtc use:RERed.FromEta. Qed.
+
+End REReds.
+
 Module LoRed.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
   (****************** Beta ***********************)
@@ -1493,6 +1562,11 @@ Module LoRed.
     move => h. elim : n a b /h=>//=.
   Qed.
 
+  Lemma ToRRed n (a b : PTm n) :
+    LoRed.R a b ->
+    RRed.R a b.
+  Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
+
 End LoRed.
 
 Module LoReds.
@@ -1547,7 +1621,7 @@ Module LoReds.
 
   Local Ltac triv := simpl in *; itauto.
 
-  Lemma FromSN : forall n,
+  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).
@@ -1571,6 +1645,12 @@ Module LoReds.
       hauto lq:on ctrs:LoRed.R.
   Qed.
 
+  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.
+  Proof. induction 1; hauto lq:on ctrs:rtc use:LoRed.ToRRed. Qed.
+
 End LoReds.
 
 
@@ -1686,6 +1766,34 @@ Lemma red_confluence n (a b c : PTm n) :
   eauto using RPar.diamond.
 Qed.
 
+Lemma red_uniquenf n (a b c : PTm n) :
+  rtc RRed.R a b ->
+  rtc RRed.R a c ->
+  nf b ->
+  nf c ->
+  b = c.
+Proof.
+  move : red_confluence; repeat move/[apply].
+  move => [d [h0 h1]].
+  move => *.
+  suff [] : b = d /\ c = d by congruence.
+  sfirstorder use:RReds.nf_refl.
+Qed.
+
+Module NeEPars.
+  Lemma R_nonelim_nf n (a b : PTm n) :
+    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).
+  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) :
   SN a ->
   rtc RERed.R a c ->
@@ -1710,7 +1818,30 @@ Lemma rered_confluence n (a b c : PTm n) :
   exists d, rtc RERed.R b d  /\ rtc RERed.R c d.
 Proof.
   move => hP hb hc.
+  have [] : SN b /\ SN c by qauto use:REReds.sn_preservation.
+  move => /LoReds.FromSN [bv [/LoReds.ToRReds /REReds.FromRReds hbv  hbv']].
+  move => /LoReds.FromSN [cv [/LoReds.ToRReds /REReds.FromRReds hcv  hcv']].
+  have [] : SN b /\ SN c by sfirstorder use:REReds.sn_preservation.
+  move : rered_standardization hbv; repeat move/[apply]. move => [bv' [hb0 hb1]].
+  move : rered_standardization hcv; repeat move/[apply]. move => [cv' [hc0 hc1]].
 
+  have [] : rtc RERed.R a bv' /\ rtc RERed.R a cv'
+    by sfirstorder use:@relations.rtc_transitive, REReds.FromRReds.
+  move : rered_standardization (hP). repeat move/[apply]. move => [bv'' [hb3 hb4]].
+  move : rered_standardization (hP). repeat move/[apply]. move => [cv'' [hc3 hc4]].
+  have hb2 : rtc NeEPar.R_nonelim bv'' bv by hauto lq:on use:@relations.rtc_transitive.
+  have hc2 : rtc NeEPar.R_nonelim cv'' cv by hauto lq:on use:@relations.rtc_transitive.
+  have [hc5 hb5] : nf cv'' /\ nf bv'' by sfirstorder use:NeEPars.R_nonelim_nf.
+  have ? : bv'' = cv'' by sfirstorder use:red_uniquenf. subst.
+  apply NeEPars.ToEReds in hb2, hc2.
+  move : ered_confluence (hb2) (hc2); repeat move/[apply].
+  move => [v [hv hv']].
+  exists v. split.
+  move /NeEPars.ToEReds /REReds.FromEReds : hb1.
+  move /REReds.FromRReds : hb0. move /REReds.FromEReds : hv. eauto using relations.rtc_transitive.
+  move /NeEPars.ToEReds /REReds.FromEReds : hc1.
+  move /REReds.FromRReds : hc0. move /REReds.FromEReds : hv'. eauto using relations.rtc_transitive.
+Qed.
 
 (* "Declarative" Joinability *)
 Module DJoin.

From 84cd0715c7a72d142a10e0e4bbde75691f628b6c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 3 Feb 2025 14:55:31 -0500
Subject: [PATCH 017/210] Prove structural properties of DJoin

---
 theories/fp_red.v | 15 +++++++++++++++
 1 file changed, 15 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 337cd11..670ce12 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1846,4 +1846,19 @@ Qed.
 (* "Declarative" Joinability *)
 Module DJoin.
   Definition R {n} (a b : PTm n) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.
+
+  Lemma refl n (a : PTm n) : R a a.
+  Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
+
+  Lemma symmetric n (a b : PTm n) : 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.
+  Proof.
+    rewrite /R.
+    move => + [ab [ha +]] [bc [+ hc]].
+    move : rered_confluence; repeat move/[apply].
+    move => [v [h0 h1]].
+    exists v. sfirstorder use:@relations.rtc_transitive.
+  Qed.
 End DJoin.

From bd7af7b297a8db021244bcdc9a4ae911aeec2dfe Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 3 Feb 2025 22:47:49 -0500
Subject: [PATCH 018/210] Add README file

---
 README.org        | 42 +++++++++++++++++++++++++++++++++++++
 theories/fp_red.v | 53 +++++++++++++++++++++++++++++++++++++++++++++++
 2 files changed, 95 insertions(+)
 create mode 100644 README.org

diff --git a/README.org b/README.org
new file mode 100644
index 0000000..69c5993
--- /dev/null
+++ b/README.org
@@ -0,0 +1,42 @@
+* Church Rosser, surjective pairing, and strong normalization
+This repository contains a mechanized proof that the lambda calculus
+with beta and eta rules for functions and pairs is in fact confluent
+for the subset of terms that are "strongly $\beta$-normalizing".
+
+Our notion of $\beta$-strong normalization, based on Abel's POPLMark
+Reloaded survey, is inductively defined
+and captures a strict subset of the usual notion of strong
+normalization in the sense that ill-formed terms such as $\pi_1
+\lambda x . x$ are rejected.
+
+* Joinability: A transitive equational relation
+The joinability relation $a \Leftrightarrow b$, which is true exactly
+when $\exists c, a \leadsto_{\beta\eta} c \wedge b
+\leadsto_{\beta\eta} c$, is transitive on the set of strongly
+normalizing terms thanks to the Church-Rosser theorem proven in this
+repository.
+
+** Joinability and logical predicates
+
+When building a logical predicate where adequacy ensures beta strong
+normalization, we can then show that two types $A \Leftrightarrow B$
+share the same meaning if both are semantically well-formed. The
+strong normalization constraint can be extracted from adequacy so we
+have transitivity of $A \Leftrightarrow B$ available in the proof that
+the logical predicate is functional over joinable terms.
+
+** Joinability and typed equality
+
+For systems with a type-directed equality, the same joinability
+relation can be used to model the equality. The fundamental theorem
+for the judgmental equality would take the following form: $\vdash a
+\equiv b \in A$ implies $\vDash a, b \in A$ and $a \Leftrightarrow b$.
+
+Note that such a result does not immediately give us the decidability
+of type conversion because the algorithm of beta-eta normalization
+may identify more terms when eta is not type-preserving. However, I
+believe Coquand's algorithm can be proven correct using the method
+described in [[https://www.researchgate.net/publication/226663076_Justifying_Algorithms_for_be-Conversion][Goguen 2005]]. We have all the ingredients needed:
+- Strong normalization of beta (needed for the termination metric)
+- Church-Rosser for $\beta$-SN terms (the disciplined expansion of
+  Coquand's algorithm preserves both SN and typing)
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 670ce12..062f9a6 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1434,6 +1434,7 @@ Module EReds.
       apply PairCong; eauto using ProjCong.
       apply ERed.PairEta.
   Qed.
+
 End EReds.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -1519,6 +1520,35 @@ Module REReds.
     rtc RERed.R a b.
   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.
+
+  #[local]Ltac solve_s :=
+    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
+
+  Lemma AbsCong n (a b : PTm (S n)) :
+    rtc RERed.R a b ->
+    rtc RERed.R (PAbs a) (PAbs b).
+  Proof. solve_s. Qed.
+
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+    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) :
+    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) :
+    rtc RERed.R a0 a1 ->
+    rtc RERed.R (PProj p a0) (PProj p a1).
+  Proof. solve_s. Qed.
+
 End REReds.
 
 Module LoRed.
@@ -1861,4 +1891,27 @@ Module DJoin.
     move => [v [h0 h1]].
     exists v. sfirstorder use:@relations.rtc_transitive.
   Qed.
+
+  Lemma AbsCong n (a b : PTm (S n)) :
+    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 n) :
+    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) :
+    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) :
+    R a0 a1 ->
+    R (PProj p a0) (PProj p a1).
+  Proof. hauto q:on use:REReds.ProjCong. Qed.
+
 End DJoin.

From d9b5ef126789c22162a34f26e726bb6cb65217f8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 4 Feb 2025 15:06:17 -0500
Subject: [PATCH 019/210] Refactor the impossible case proof

---
 syntax.sig                   |  11 +-
 theories/Autosubst2/syntax.v | 117 +++++++++---
 theories/fp_red.v            | 344 ++++++++++++++++++++---------------
 3 files changed, 294 insertions(+), 178 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index e2bafac..e17d7ea 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -1,15 +1,18 @@
 PTm(VarPTm) : Type
 PTag : Type
-Ty : Type
+BTag : Type
 
-Fun : Ty -> Ty -> Ty
-Prod : Ty -> Ty -> Ty
-Void : Ty
+nat : Type
 
 PL : PTag
 PR : PTag
 
+PPi : BTag
+PSig : BTag
+
 PAbs : (bind PTm in PTm) -> PTm
 PApp : PTm -> PTm -> PTm
 PPair : PTm -> PTm -> PTm
 PProj : PTag -> PTm -> PTm
+PBind : BTag -> PTm -> (bind PTm in PTm) -> PTm
+PUniv : nat -> PTm
\ No newline at end of file
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index 78c4fec..fbdf45d 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -5,6 +5,20 @@ Require Import Setoid Morphisms Relation_Definitions.
 
 Module Core.
 
+Inductive BTag : Type :=
+  | PPi : BTag
+  | PSig : BTag.
+
+Lemma congr_PPi : PPi = PPi.
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma congr_PSig : PSig = PSig.
+Proof.
+exact (eq_refl).
+Qed.
+
 Inductive PTag : Type :=
   | PL : PTag
   | PR : PTag.
@@ -24,7 +38,9 @@ Inductive PTm (n_PTm : nat) : Type :=
   | 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.
+  | 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.
 
 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.
@@ -56,6 +72,23 @@ exact (eq_trans (eq_trans eq_refl (ap (fun x => PProj m_PTm x s1) H0))
          (ap (fun x => PProj m_PTm t0 x) H1)).
 Qed.
 
+Lemma congr_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.
+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)).
+Qed.
+
+Lemma congr_PUniv {m_PTm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
+  PUniv m_PTm s0 = PUniv m_PTm t0.
+Proof.
+exact (eq_trans eq_refl (ap (fun x => PUniv m_PTm x) H0)).
+Qed.
+
 Lemma upRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n) :
   fin (S m) -> fin (S n).
 Proof.
@@ -76,6 +109,9 @@ Fixpoint ren_PTm {m_PTm : nat} {n_PTm : nat}
   | 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
   end.
 
 Lemma up_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm) :
@@ -102,6 +138,10 @@ Fixpoint subst_PTm {m_PTm : nat} {n_PTm : nat}
   | 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
   end.
 
 Lemma upId_PTm_PTm {m_PTm : nat} (sigma : fin m_PTm -> PTm m_PTm)
@@ -140,6 +180,10 @@ Fixpoint idSubst_PTm {m_PTm : nat} (sigma_PTm : fin m_PTm -> PTm m_PTm)
         (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 =>
+      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)
   end.
 
 Lemma upExtRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n)
@@ -180,6 +224,11 @@ ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
         (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0) (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma upExt_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm)
@@ -221,6 +270,11 @@ subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
         (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0) (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma up_ren_ren_PTm_PTm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
@@ -262,6 +316,12 @@ Fixpoint compRenRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0)
         (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma up_ren_subst_PTm_PTm {k : nat} {l : nat} {m_PTm : nat}
@@ -312,6 +372,12 @@ Fixpoint compRenSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0)
         (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma up_subst_ren_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
@@ -382,6 +448,12 @@ ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0)
         (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma up_subst_subst_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
@@ -454,6 +526,12 @@ subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0)
         (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma renRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
@@ -566,6 +644,11 @@ Fixpoint rinst_inst_PTm {m_PTm : nat} {n_PTm : nat}
         (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
   | PProj _ s0 s1 =>
       congr_PProj (eq_refl s0) (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
+  | 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)
   end.
 
 Lemma rinstInst'_PTm {m_PTm : nat} {n_PTm : nat}
@@ -637,30 +720,6 @@ Proof.
 exact (fun x => eq_refl).
 Qed.
 
-Inductive Ty : Type :=
-  | Fun : Ty -> Ty -> Ty
-  | Prod : Ty -> Ty -> Ty
-  | Void : Ty.
-
-Lemma congr_Fun {s0 : Ty} {s1 : Ty} {t0 : Ty} {t1 : Ty} (H0 : s0 = t0)
-  (H1 : s1 = t1) : Fun s0 s1 = Fun t0 t1.
-Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => Fun x s1) H0))
-         (ap (fun x => Fun t0 x) H1)).
-Qed.
-
-Lemma congr_Prod {s0 : Ty} {s1 : Ty} {t0 : Ty} {t1 : Ty} (H0 : s0 = t0)
-  (H1 : s1 = t1) : Prod s0 s1 = Prod t0 t1.
-Proof.
-exact (eq_trans (eq_trans eq_refl (ap (fun x => Prod x s1) H0))
-         (ap (fun x => Prod t0 x) H1)).
-Qed.
-
-Lemma congr_Void : Void = Void.
-Proof.
-exact (eq_refl).
-Qed.
-
 Class Up_PTm X Y :=
     up_PTm : X -> Y.
 
@@ -796,6 +855,10 @@ Core.
 
 Arguments VarPTm {n_PTm}.
 
+Arguments PUniv {n_PTm}.
+
+Arguments PBind {n_PTm}.
+
 Arguments PProj {n_PTm}.
 
 Arguments PPair {n_PTm}.
@@ -804,9 +867,9 @@ Arguments PApp {n_PTm}.
 
 Arguments PAbs {n_PTm}.
 
-#[global] Hint Opaque subst_PTm: rewrite.
+#[global]Hint Opaque subst_PTm: rewrite.
 
-#[global] Hint Opaque ren_PTm: rewrite.
+#[global]Hint Opaque ren_PTm: rewrite.
 
 End Extra.
 
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 062f9a6..1e809c5 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -47,7 +47,13 @@ Module EPar.
     R a0 a1 ->
     R (PProj p a0) (PProj p a1)
   | VarTm i :
-    R (VarPTm i) (VarPTm i).
+    R (VarPTm i) (VarPTm i)
+  | Univ i :
+    R (PUniv i) (PUniv i)
+  | BindCong p A0 A1 B0 B1 :
+    R A0 A1 ->
+    R B0 B1 ->
+    R (PBind p A0 B0) (PBind p A1 B1).
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -126,6 +132,12 @@ with SN  {n} : PTm n -> Prop :=
   TRedSN a b ->
   SN b ->
   SN a
+| N_Bind p A B :
+  SN A ->
+  SN B ->
+  SN (PBind p A B)
+| N_Univ i :
+  SN (PUniv i)
 with TRedSN {n} : PTm n -> PTm n -> Prop :=
 | N_β a b :
   SN b ->
@@ -146,6 +158,63 @@ with TRedSN {n} : PTm n -> PTm n -> Prop :=
 
 Derive Dependent Inversion tred_inv with (forall n (a b : PTm n), TRedSN a b) Sort Prop.
 
+Definition ishf {n} (a : PTm n) :=
+  match a with
+  | PPair _ _ => true
+  | PAbs _ => true
+  | PUniv _ => true
+  | PBind _ _ _ => true
+  | _ => false
+  end.
+
+Definition ispair {n} (a : PTm n) :=
+  match a with
+  | PPair _ _ => true
+  | _ => false
+  end.
+
+Definition isabs {n} (a : PTm n) :=
+  match a with
+  | PAbs _ => true
+  | _ => false
+  end.
+
+Definition ishf_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  ishf (ren_PTm ξ a) = ishf a.
+Proof. case : a => //=. Qed.
+
+Definition isabs_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  isabs (ren_PTm ξ a) = isabs a.
+Proof. case : a => //=. Qed.
+
+Definition ispair_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  ispair (ren_PTm ξ a) = ispair a.
+Proof. case : a => //=. Qed.
+
+
+Lemma PProj_imp n p a :
+  @ishf n 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 => //=.
+  hauto lq:on inv:SNe,PTm.
+  hauto lq:on inv:TRedSN.
+Qed.
+
+Lemma PAbs_imp n a b :
+  @ishf n 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=>//=.
+  hauto lq:on inv:SNe,PTm.
+  hauto lq:on inv:TRedSN.
+Qed.
+
 Lemma PProjAbs_imp n p (a : PTm (S n)) :
   ~ SN (PProj p (PAbs a)).
 Proof.
@@ -156,7 +225,7 @@ Proof.
   hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma PProjPair_imp n (a b0 b1 : PTm n ) :
+Lemma PAppPair_imp n (a b0 b1 : PTm n ) :
   ~ SN (PApp (PPair b0 b1) a).
 Proof.
   move E : (PApp (PPair b0 b1) a) => u hu.
@@ -166,6 +235,26 @@ Proof.
   hauto lq:on inv:TRedSN.
 Qed.
 
+Lemma PAppBind_imp n p (A : PTm n) B b :
+  ~ SN (PApp (PBind p A B) b).
+Proof.
+  move E :(PApp (PBind p A B) b) => u hu.
+  move : p A B b E.
+  elim : n u /hu=> //=.
+  hauto lq:on inv:SNe.
+  hauto lq:on inv:TRedSN.
+Qed.
+
+Lemma PProjBind_imp n p p' (A : PTm n) 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=>//=.
+  hauto lq:on inv:SNe.
+  hauto lq:on inv:TRedSN.
+Qed.
+
 Scheme sne_ind := Induction for SNe Sort Prop
   with sn_ind := Induction for SN Sort Prop
   with sred_ind := Induction for TRedSN Sort Prop.
@@ -179,6 +268,8 @@ Fixpoint ne {n} (a : PTm n) :=
   | PAbs a => false
   | PPair _ _ => false
   | PProj _ a => ne a
+  | PUniv _ => false
+  | PBind _ _ _ => false
   end
 with nf {n} (a : PTm n) :=
   match a with
@@ -187,6 +278,8 @@ with nf {n} (a : PTm n) :=
   | PAbs a => nf a
   | PPair a b => nf a && nf b
   | PProj _ a => ne a
+  | PUniv _ => true
+  | PBind _ A B => nf A && nf B
 end.
 
 Lemma ne_nf n a : @ne n a -> nf a.
@@ -350,6 +443,8 @@ Proof.
     + sauto lq:on.
   - sauto lq:on.
   - sauto lq:on.
+  - sauto lq:on.
+  - sauto lq:on.
   - move => a b ha iha c h0.
     inversion h0; subst.
     inversion H1; subst.
@@ -410,7 +505,13 @@ Module RRed.
     R (PPair a b0) (PPair a b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (PProj p a0) (PProj p a1).
+    R (PProj p a0) (PProj p a1)
+  | BindCong0 p A0 A1 B :
+    R A0 A1 ->
+    R (PBind p A0 B) (PBind p A1 B)
+  | BindCong1 p A B0 B1 :
+    R B0 B1 ->
+    R (PBind p A B0) (PBind p A B1).
 
   Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
@@ -488,7 +589,13 @@ Module RPar.
     R a0 a1 ->
     R (PProj p a0) (PProj p a1)
   | VarTm i :
-    R (VarPTm i) (VarPTm i).
+    R (VarPTm i) (VarPTm i)
+  | Univ i :
+    R (PUniv i) (PUniv i)
+  | BindCong p A0 A1 B0 B1 :
+    R A0 A1 ->
+    R B0 B1 ->
+    R (PBind p A0 B0) (PBind p A1 B1).
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -581,6 +688,8 @@ Module RPar.
     | PPair a b => PPair (tstar a) (tstar b)
     | PProj p (PPair a b) => if p is PL then (tstar a) else (tstar b)
     | PProj p a => PProj p (tstar a)
+    | PUniv i => PUniv i
+    | PBind p A B => PBind p (tstar A) (tstar B)
     end.
 
   Lemma triangle n (a b : PTm n) :
@@ -609,6 +718,8 @@ Module RPar.
         elim /inv : ha => //=_ > ? ? [*]. subst.
         apply : ProjPair'; eauto using refl.
     - hauto lq:on ctrs:R inv:R.
+    - hauto lq:on ctrs:R inv:R.
+    - hauto lq:on ctrs:R inv:R.
   Qed.
 
   Lemma diamond n (a b c : PTm n) :
@@ -629,6 +740,8 @@ Proof.
   - hauto lq:on ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN.
   - hauto q:on ctrs:SN inv:SN, TRedSN'.
+  - hauto lq:on ctrs:SN inv:RPar.R.
+  - hauto lq:on ctrs:SN inv:RPar.R.
   - move => a b ha iha hb ihb.
     elim /RPar.inv : ihb => //=_.
     + move => a0 a1 b0 b1 ha0 hb0 [*]. subst.
@@ -650,91 +763,6 @@ Proof.
   - sauto.
 Qed.
 
-Module EPar'.
-  Inductive R {n} : PTm n -> PTm n ->  Prop :=
-  (****************** Eta ***********************)
-  | AppEta a :
-    R (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) a
-  | PairEta a :
-    R (PPair (PProj PL a) (PProj PR a)) a
-  (*************** Congruence ********************)
-  | AbsCong a0 a1 :
-    R a0 a1 ->
-    R (PAbs a0) (PAbs a1)
-  | AppCong0 a0 a1 b :
-    R a0 a1 ->
-    R (PApp a0 b) (PApp a1 b)
-  | AppCong1 a b0 b1 :
-    R b0 b1 ->
-    R (PApp a b0) (PApp a b1)
-  | PairCong0 a0 a1 b :
-    R a0 a1 ->
-    R (PPair a0 b) (PPair a1 b)
-  | PairCong1 a b0 b1 :
-    R b0 b1 ->
-    R (PPair a b0) (PPair a b1)
-  | ProjCong p a0 a1 :
-    R a0 a1 ->
-    R (PProj p a0) (PProj p a1).
-
-  Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
-
-  Lemma AppEta' n a (u : PTm n) :
-    u = (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) ->
-    R u a.
-  Proof. move => ->. apply AppEta. Qed.
-
-  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
-    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
-  Proof.
-    move => h. move : m ξ.
-    elim : n a b /h.
-
-    move => n a m ξ /=.
-    eapply AppEta'; eauto. by asimpl.
-    all : qauto ctrs:R.
-  Qed.
-
-  Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
-    (forall i, R (ρ0 i) (ρ1 i)) ->
-    (forall i, R ((funcomp (ren_PTm ξ) ρ0) i) ((funcomp (ren_PTm ξ) ρ1) i)).
-  Proof. eauto using renaming. Qed.
-
-End EPar'.
-
-Module EPars.
-
-  #[local]Ltac solve_s_rec :=
-  move => *; eapply rtc_l; eauto;
-         hauto lq:on ctrs:EPar'.R.
-
-  #[local]Ltac solve_s :=
-    repeat (induction 1; last by solve_s_rec); apply rtc_refl.
-
-  Lemma AbsCong n (a b : PTm (S n)) :
-    rtc EPar'.R a b ->
-    rtc EPar'.R (PAbs a) (PAbs b).
-  Proof. solve_s. Qed.
-
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
-    rtc EPar'.R a0 a1 ->
-    rtc EPar'.R b0 b1 ->
-    rtc EPar'.R (PApp a0 b0) (PApp a1 b1).
-  Proof. solve_s. Qed.
-
-  Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
-    rtc EPar'.R a0 a1 ->
-    rtc EPar'.R b0 b1 ->
-    rtc EPar'.R (PPair a0 b0) (PPair a1 b1).
-  Proof. solve_s. Qed.
-
-  Lemma ProjCong n p (a0 a1  : PTm n) :
-    rtc EPar'.R a0 a1 ->
-    rtc EPar'.R (PProj p a0) (PProj p a1).
-  Proof. solve_s. Qed.
-
-End EPars.
-
 Module RReds.
 
   #[local]Ltac solve_s_rec :=
@@ -766,6 +794,12 @@ Module RReds.
     rtc RRed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
+  Lemma BindCong n p (A0 A1 : PTm n) 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) :
     rtc RRed.R a b -> rtc RRed.R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
@@ -775,7 +809,7 @@ Module RReds.
   Lemma FromRPar n (a b : PTm n) (h : RPar.R a b) :
     rtc RRed.R a b.
   Proof.
-    elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl.
+    elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong.
     move => n a0 a1 b0 b1 ha iha hb ihb.
     apply : rtc_r; last by apply RRed.AppAbs.
     by eauto using AppCong, AbsCong.
@@ -806,19 +840,6 @@ Proof.
   move : m ξ. elim : n / a => //=; solve [hauto b:on].
 Qed.
 
-Lemma ne_epar n (a b : PTm n) (h : EPar'.R a b ) :
-  (ne a -> ne b) /\ (nf a -> nf b).
-Proof.
-  elim : n a b /h=>//=; hauto qb:on use:ne_nf_ren, ne_nf.
-Qed.
-
-Definition ishf {n} (a : PTm n) :=
-  match a with
-  | PPair _ _ => true
-  | PAbs _ => true
-  | _ => false
-  end.
-
 Module NeEPar.
   Inductive R_nonelim {n} : PTm n -> PTm n ->  Prop :=
   (****************** Eta ***********************)
@@ -847,6 +868,12 @@ Module NeEPar.
     R_nonelim (PProj p a0) (PProj p a1)
   | VarTm i :
     R_nonelim (VarPTm i) (VarPTm i)
+  | Univ i :
+    R_nonelim (PUniv i) (PUniv i)
+  | BindCong p A0 A1 B0 B1 :
+    R_nonelim A0 A1 ->
+    R_nonelim B0 B1 ->
+    R_nonelim (PBind p A0 B0) (PBind p A1 B1)
   with R_elim {n} : PTm n -> PTm n -> Prop :=
   | NAbsCong a0 a1 :
     R_nonelim a0 a1 ->
@@ -863,7 +890,13 @@ Module NeEPar.
     R_elim a0 a1 ->
     R_elim (PProj p a0) (PProj p a1)
   | NVarTm i :
-    R_elim (VarPTm i) (VarPTm i).
+    R_elim (VarPTm i) (VarPTm i)
+  | NUniv i :
+    R_elim (PUniv i) (PUniv i)
+  | NBindCong p A0 A1 B0 B1 :
+    R_nonelim A0 A1 ->
+    R_nonelim B0 B1 ->
+    R_elim (PBind p A0 B0) (PBind p A1 B1).
 
   Scheme epar_elim_ind := Induction for R_elim Sort Prop
       with epar_nonelim_ind := Induction for R_nonelim Sort Prop.
@@ -881,6 +914,7 @@ Module NeEPar.
       qauto l:on inv:R_elim.
     - hauto lb:on.
     - hauto lq:on inv:R_elim.
+    - hauto b:on.
     - move => a0 a1 /negP ha' ha ih ha1.
       have {ih} := ih ha1.
       move => ha0.
@@ -897,6 +931,7 @@ Module NeEPar.
       move : ha h0. hauto lq:on inv:R_elim.
     - hauto lb: on drew: off.
     - hauto lq:on rew:off inv:R_elim.
+    - sfirstorder b:on.
   Qed.
 
   Lemma R_nonelim_nothf n (a b : PTm n) :
@@ -927,11 +962,17 @@ Module Type NoForbid.
 
   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_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 PAbs_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 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_BindInv : forall n p (A : PTm n) B, P (PBind p A B) -> P A /\ P B.
+
   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.
@@ -970,11 +1011,10 @@ Module SN_NoForbid <: NoForbid.
   Lemma P_RRed : forall n (a b : PTm n), RRed.R a b -> P a -> P b.
   Proof. hauto q:on use:red_sn_preservation, RPar.FromRRed. Qed.
 
-  Lemma P_AppPair : forall n (a b c : PTm n), ~ P (PApp (PPair a b) c).
-  Proof. sfirstorder use:PProjPair_imp. Qed.
-
-  Lemma P_ProjAbs : forall n p (a : PTm (S n)), ~ P (PProj p (PAbs a)).
-  Proof. sfirstorder use:PProjAbs_imp. Qed.
+  Lemma PAbs_imp : forall n a b, @ishf n a -> ~~ isabs a -> ~ P (PApp a b).
+    sfirstorder use:fp_red.PAbs_imp. Qed.
+  Lemma PProj_imp : forall  n p a, @ishf n a -> ~~ ispair a -> ~ P (PProj p a).
+    sfirstorder use:fp_red.PProj_imp. Qed.
 
   Lemma P_AppInv : forall n (a b : PTm n), P (PApp a b) -> P a /\ P b.
   Proof. sfirstorder use:SN_AppInv. Qed.
@@ -986,6 +1026,13 @@ Module SN_NoForbid <: NoForbid.
   Lemma P_ProjInv : forall n p (a : PTm n), 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.
+  Proof.
+    move => n 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.
+  Qed.
+
   Lemma P_AbsInv : forall n (a : PTm (S n)), P (PAbs a) -> P a.
   Proof.
     move => n a. move E : (PAbs a) => u h.
@@ -996,13 +1043,19 @@ Module SN_NoForbid <: NoForbid.
   Lemma P_renaming : forall n m (ξ : fin n -> fin m) 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)).
+  Proof. sfirstorder use:PProjBind_imp. Qed.
+
+  Lemma P_AppBind : forall n p (A : PTm n) B b, ~ P (PApp (PBind p A B) b).
+  Proof. sfirstorder use:PAppBind_imp. Qed.
+
 End SN_NoForbid.
 
 Module NoForbid_FactSN := NoForbid_Fact SN_NoForbid.
 
 Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
   Import M MFacts.
-  #[local]Hint Resolve P_EPar P_RRed P_AppPair P_ProjAbs : forbid.
+  #[local]Hint Resolve P_EPar P_RRed PAbs_imp PProj_imp : forbid.
 
   Lemma η_split n (a0 a1 : PTm n) :
     EPar.R a0 a1 ->
@@ -1020,9 +1073,8 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       by eauto using RReds.renaming.
       apply NeEPar.AppEta=>//.
       sfirstorder use:NeEPar.R_nonelim_nothf.
-
-      case : b ih0 ih1 => //=.
-      + move => p ih0 ih1 _.
+      case /orP : (orbN (isabs b)).
+      + case : b ih0 ih1 => //= p ih0 ih1 _ _.
         set q := PAbs _.
         suff : rtc RRed.R q (PAbs p) by sfirstorder.
         subst q.
@@ -1033,16 +1085,14 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         apply : RRed.AbsCong => /=.
         apply RRed.AppAbs'. by asimpl.
       (* violates SN *)
-      + move => p p0 h. exfalso.
-        have : P (PApp (ren_PTm shift a0) (VarPTm var_zero))
-          by sfirstorder use:P_AbsInv.
-
-        have : rtc RRed.R (PApp (ren_PTm shift a0) (VarPTm var_zero))
-                 (PApp (ren_PTm shift (PPair p p0)) (VarPTm var_zero))
-          by hauto lq:on use:RReds.AppCong, RReds.renaming, rtc_refl.
-
-        move : P_RReds. repeat move/[apply] => /=.
-        hauto l:on use:P_AppPair.
+      + move /P_AbsInv in hP.
+        have {}hP : P (PApp (ren_PTm shift b) (VarPTm var_zero))
+          by sfirstorder use:P_RReds, RReds.AppCong, @rtc_refl, RReds.renaming.
+        move => ? ?.
+        have ? : ~~ isabs (ren_PTm shift b) by scongruence use:isabs_ren.
+        have ? : ishf (ren_PTm shift b) by scongruence use:ishf_ren.
+        exfalso.
+        sfirstorder use:PAbs_imp.
     - move => n a0 a1 h ih /[dup] hP.
       move /P_PairInv => [/P_ProjInv + _].
       move : ih => /[apply].
@@ -1051,16 +1101,9 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       exists (PPair (PProj PL b) (PProj PR b)).
       split. sfirstorder use:RReds.PairCong,RReds.ProjCong.
       hauto lq:on ctrs:NeEPar.R_nonelim use:NeEPar.R_nonelim_nothf.
-
-      case : b ih0 ih1 => //=.
-      (* violates SN *)
-      + move => p ?. exfalso.
-        have {}hP : P (PProj PL a0) by sfirstorder use:P_PairInv.
-        have : rtc RRed.R (PProj PL a0) (PProj PL (PAbs p))
-          by eauto using RReds.ProjCong.
-        move : P_RReds hP. repeat move/[apply] => /=.
-        sfirstorder use:P_ProjAbs.
-      + move => t0 t1 ih0 h1 _.
+      case /orP : (orbN (ispair b)).
+      + case : b ih0 ih1 => //=.
+        move => t0 t1 ih0 h1 _ _.
         exists (PPair t0 t1).
         split => //=.
         apply RReds.PairCong.
@@ -1068,6 +1111,12 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         apply RRed.ProjPair.
         apply : rtc_r; eauto using RReds.ProjCong.
         apply RRed.ProjPair.
+      + move => ? ?. exfalso.
+        move/P_PairInv : hP=>[hP _].
+        have : rtc RRed.R (PProj PL a0) (PProj PL b)
+          by eauto using RReds.ProjCong.
+        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 => /[dup] hP /P_AppInv [hP0 hP1].
@@ -1076,8 +1125,9 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       case /orP : (orNb (ishf a2)) => [h|].
       + exists (PApp a2 b2). split; first by eauto using RReds.AppCong.
         hauto lq:on ctrs:NeEPar.R_nonelim use:NeEPar.R_nonelim_nothf.
-      + case : a2 iha0 iha1 => //=.
-        * move => p h0 h1 _.
+      + case /orP : (orbN (isabs a2)).
+        (* case : a2 iha0 iha1 => //=. *)
+        * case : a2 iha0 iha1 => //= p h0 h1 _ _.
           inversion h1; subst.
           ** exists (PApp a2 b2).
              split.
@@ -1087,11 +1137,9 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
              hauto lq:on ctrs:NeEPar.R_nonelim.
           ** hauto lq:on ctrs:NeEPar.R_nonelim,NeEPar.R_elim use:RReds.AppCong.
         (* Impossible *)
-        * move => u0 u1 h. exfalso.
-          have : rtc RRed.R (PApp a0 b0) (PApp (PPair u0 u1) b0)
-            by hauto lq:on ctrs:rtc use:RReds.AppCong.
-          move : P_RReds hP; repeat move/[apply].
-          sfirstorder use:P_AppPair.
+        * move =>*. exfalso.
+          have : P (PApp a2 b0) by sfirstorder use:RReds.AppCong, @rtc_refl, P_RReds.
+          sfirstorder use:PAbs_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 : ih => /[apply]. move => [a2 [iha0 iha1]].
@@ -1100,13 +1148,13 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       split.  eauto using RReds.ProjCong.
       qauto l:on ctrs:NeEPar.R_nonelim, NeEPar.R_elim use:NeEPar.R_nonelim_nothf.
 
-      case : a2 iha0 iha1 => //=.
-      + move => u iha0. exfalso.
-        have : rtc RRed.R (PProj p a0) (PProj p (PAbs u))
+      case /orP : (orNb (ispair a2)).
+      + move => *. exfalso.
+        have : rtc RRed.R (PProj p a0) (PProj p a2)
           by sfirstorder use:RReds.ProjCong ctrs:rtc.
         move : P_RReds hP. repeat move/[apply].
-        sfirstorder use:P_ProjAbs.
-      + move => u0 u1 iha0 iha1 _.
+        sfirstorder use:PProj_imp.
+      + case : a2 iha0 iha1 => //= u0 u1 iha0 iha1 _ _.
         inversion iha1; subst.
         * exists (PProj p a2). split.
           apply : rtc_r.
@@ -1115,6 +1163,8 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           hauto lq:on ctrs:NeEPar.R_nonelim.
         * hauto lq:on ctrs:NeEPar.R_nonelim,NeEPar.R_elim use:RReds.ProjCong.
     - hauto lq:on ctrs:rtc, NeEPar.R_nonelim.
+    - hauto l:on.
+    - hauto lq:on ctrs:NeEPar.R_nonelim, rtc use:RReds.BindCong, P_BindInv.
   Qed.
 
 

From e923194e3d2a06b3aeba11283a6ae34eb65c87bf Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 4 Feb 2025 15:29:47 -0500
Subject: [PATCH 020/210] Finish adding pi and sigma types

---
 theories/fp_red.v | 84 ++++++++++++++++++++++++++++++++++++++++-------
 1 file changed, 73 insertions(+), 11 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 1e809c5..ec43c8f 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -911,7 +911,7 @@ Module NeEPar.
     - 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.
-      qauto l:on inv:R_elim.
+      hauto q:on inv:R_elim.
     - hauto lb:on.
     - hauto lq:on inv:R_elim.
     - hauto b:on.
@@ -1218,7 +1218,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           have : rtc RRed.R (PApp a0 b0) (PApp (PPair (PProj PL a1) (PProj PR a1)) b0)
             by qauto l:on ctrs:rtc use:RReds.AppCong.
           move : P_RReds hP. repeat move/[apply].
-          sfirstorder use:P_AppPair.
+          sfirstorder use:PAbs_imp.
         * exists (subst_PTm (scons b0 VarPTm) a1).
           split.
           apply : rtc_r; last by apply RRed.AppAbs.
@@ -1245,7 +1245,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         move : η_split hP'  ha; repeat move/[apply].
         move => [a1 [h0 h1]].
         inversion h1; subst.
-        * qauto l:on ctrs:rtc use:RReds.ProjCong, P_ProjAbs, P_RReds.
+        * sauto q:on ctrs:rtc use:RReds.ProjCong, PProj_imp, P_RReds.
         * inversion H0; subst.
           exists (if p is PL then a1 else b1).
           split; last by scongruence use:NeEPar.ToEPar.
@@ -1270,6 +1270,8 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         move : iha hP' h0;repeat move/[apply].
         hauto lq:on ctrs:rtc, EPar.R use:RReds.ProjCong.
     - hauto lq:on inv:RRed.R.
+    - hauto lq:on inv:RRed.R ctrs:rtc.
+    - sauto lq:on ctrs:EPar.R, rtc use:RReds.BindCong, P_BindInv, @relations.rtc_transitive.
   Qed.
 
   Lemma η_postponement_star n a b c :
@@ -1306,7 +1308,6 @@ End UniqueNF.
 
 Module SN_UniqueNF := UniqueNF SN_NoForbid NoForbid_FactSN.
 
-
 Module ERed.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
 
@@ -1334,7 +1335,13 @@ Module ERed.
     R (PPair a b0) (PPair a b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (PProj p a0) (PProj p a1).
+    R (PProj p a0) (PProj p a1)
+  | BindCong0 p A0 A1 B :
+    R A0 A1 ->
+    R (PBind p A0 B) (PBind p A1 B)
+  | BindCong1 p A B0 B1 :
+    R B0 B1 ->
+    R (PBind p A B0) (PBind p A B1).
 
   Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
@@ -1378,6 +1385,13 @@ Module ERed.
     apply iha in h. by subst.
     destruct i, j=>//=.
     hauto l:on.
+
+    move => n p A ihA B ihB m ξ []//=.
+    move => b A0 B0  hξ [?]. subst.
+    move => ?. have ? : A0 = A by firstorder. subst.
+    move => ?. have : B = B0. apply : ihB; eauto.
+    sauto.
+    congruence.
   Qed.
 
   Lemma AppEta' n a u :
@@ -1430,9 +1444,14 @@ Module ERed.
       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ξ.
+      case => //= p' A2 B2 [*]. subst.
+      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.
 
-
 End ERed.
 
 Module EReds.
@@ -1466,6 +1485,13 @@ Module EReds.
     rtc ERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
+  Lemma BindCong n p (A0 A1 : PTm n) 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 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).
   Proof. induction 1; hauto l:on use:ERed.renaming ctrs:rtc. Qed.
@@ -1474,7 +1500,7 @@ Module EReds.
     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.
+    move => h. elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong.
     - move => n 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.
@@ -1523,7 +1549,13 @@ Module RERed.
     R (PPair a b0) (PPair a b1)
   | ProjCong p a0 a1 :
     R a0 a1 ->
-    R (PProj p a0) (PProj p a1).
+    R (PProj p a0) (PProj p a1)
+  | BindCong0 p A0 A1 B :
+    R A0 A1 ->
+    R (PBind p A0 B) (PBind p A1 B)
+  | BindCong1 p A B0 B1 :
+    R B0 B1 ->
+    R (PBind p A B0) (PBind p A B1).
 
   Lemma ToBetaEta n (a b : PTm n) :
     R a b ->
@@ -1599,6 +1631,12 @@ Module REReds.
     rtc RERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
+  Lemma BindCong n p (A0 A1 : PTm n) 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.
+
 End REReds.
 
 Module LoRed.
@@ -1632,7 +1670,14 @@ Module LoRed.
   | ProjCong p a0 a1 :
     ~~ ishf a0 ->
     R a0 a1 ->
-    R (PProj p a0) (PProj p a1).
+    R (PProj p a0) (PProj p a1)
+  | BindCong0 p A0 A1 B :
+    R A0 A1 ->
+    R (PBind p A0 B) (PBind p A1 B)
+  | BindCong1 p A B0 B1 :
+    nf A ->
+    R B0 B1 ->
+    R (PBind p A B0) (PBind p A B1).
 
   Lemma hf_preservation n (a b : PTm n) :
     LoRed.R a b ->
@@ -1699,6 +1744,13 @@ Module LoReds.
     rtc LoRed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
+  Lemma BindCong n p (A0 A1 : PTm n) 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.
+
   Local Ltac triv := simpl in *; itauto.
 
   Lemma FromSN_mutual : forall n,
@@ -1714,6 +1766,8 @@ Module LoReds.
     - hauto q:on use:LoReds.AbsCong solve+:triv.
     - sfirstorder use:ne_nf.
     - hauto lq:on ctrs:rtc.
+    - hauto lq:on use:LoReds.BindCong solve+:triv.
+    - hauto lq:on ctrs:rtc.
     - qauto ctrs:LoRed.R.
     - move => n a0 a1 b hb ihb h.
       have : ~~ ishf a0 by inversion h.
@@ -1737,10 +1791,12 @@ End LoReds.
 Fixpoint size_PTm {n} (a : PTm n) :=
   match a with
   | VarPTm _ => 1
-  | PAbs a => 1 + size_PTm a
+  | 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 => 1 + Nat.add (size_PTm a) (size_PTm b)
+  | PPair a b => 3 + Nat.add (size_PTm a) (size_PTm b)
+  | PUniv _ => 3
+  | PBind p A B => 3 + Nat.add (size_PTm A) (size_PTm B)
   end.
 
 Lemma size_PTm_ren n m (ξ : fin n -> fin m) a : size_PTm (ren_PTm ξ a) = size_PTm a.
@@ -1825,6 +1881,12 @@ Proof.
     + hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
     + hauto lq:on ctrs:rtc use:EReds.PairCong.
   - qauto l:on inv:ERed.R use:EReds.ProjCong.
+  - move => p A0 A1 B hA ihA u.
+    elim /ERed.inv => //=_;
+      hauto lq:on ctrs:rtc use:EReds.BindCong.
+  - move => p A B0 B1 hB ihB u.
+    elim /ERed.inv => //=_;
+      hauto lq:on ctrs:rtc use:EReds.BindCong.
 Qed.
 
 Lemma ered_confluence n (a b c : PTm n) :

From ee24f8093ecd4f4a2759aaabdb9e231c41bf635b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 4 Feb 2025 16:05:02 -0500
Subject: [PATCH 021/210] Add logical relation for SN

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

diff --git a/theories/logrel.v b/theories/logrel.v
new file mode 100644
index 0000000..1e638e6
--- /dev/null
+++ b/theories/logrel.v
@@ -0,0 +1,140 @@
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+Require Import fp_red.
+From Hammer Require Import Tactics.
+From Equations Require Import Equations.
+Require Import ssreflect ssrbool.
+Require Import Logic.PropExtensionality (propositional_extensionality).
+From stdpp Require Import relations (rtc(..), rtc_subrel).
+Import Psatz.
+
+Definition ProdSpace {n} (PA : PTm n -> Prop)
+  (PF : PTm n -> (PTm n -> 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 :=
+  SNe t \/ 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.
+
+Reserved Notation "⟦ A ⟧ i ;; I ↘ S" (at level 70).
+
+Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop) -> Prop :=
+| InterpExt_Ne A :
+  SNe A ->
+  ⟦ A ⟧ i ;; I ↘ SNe
+| InterpExt_Bind p A B PA PF :
+  ⟦ A ⟧ i ;; I ↘ PA ->
+  (forall a, PA a -> exists PB, PF a PB) ->
+  (forall a PB, PF a PB -> ⟦ subst_PTm (scons a VarPTm) B ⟧ i ;; I ↘ PB) ->
+  ⟦ PBind p A B ⟧ i ;; I ↘ BindSpace p PA PF
+
+| InterpExt_Univ j :
+  j < i ->
+  ⟦ PUniv j ⟧ i ;; I ↘ (I j)
+
+| InterpExt_Step A A0 PA :
+  TRedSN A A0 ->
+  ⟦ A0 ⟧ i ;; I ↘ PA ->
+  ⟦ 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) :
+  PF = I j ->
+  j < i ->
+  ⟦ PUniv j ⟧ i ;; I ↘ PF.
+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
+                     (fun j A =>
+                        match j <? i  with
+                        | left _ => exists PA, InterpUnivN n j A PA
+                        | right _ => False
+                        end).
+Arguments InterpUnivN {n}.
+
+Lemma InterpExt_lt_impl n i I I' A (PA : PTm n -> Prop) :
+  (forall j, j < i -> I j = I' j) ->
+  ⟦ A ⟧ i ;; I ↘ PA ->
+  ⟦ A ⟧ i ;; I' ↘ PA.
+Proof.
+  move => hI h.
+  elim : A PA /h.
+  - hauto q:on ctrs:InterpExt.
+  - hauto lq:on rew:off ctrs:InterpExt.
+  - hauto q:on ctrs:InterpExt.
+  - hauto lq:on ctrs:InterpExt.
+Qed.
+
+Lemma InterpExt_lt_eq n i I I' A (PA : PTm n -> Prop) :
+  (forall j, j < i -> I j = I' j) ->
+  ⟦ A ⟧ i ;; I ↘ PA =
+  ⟦ A ⟧ i ;; I' ↘ PA.
+Proof.
+  move => hI. apply propositional_extensionality.
+  have : forall j, j < i -> I' j = I j by sfirstorder.
+  firstorder using InterpExt_lt_impl.
+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).
+Proof.
+  simp InterpUnivN.
+  extensionality A. extensionality PA.
+  set I0 := (fun _ => _).
+  set I1 := (fun _ => _).
+  apply InterpExt_lt_eq.
+  hauto q:on.
+Qed.
+
+#[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv.
+
+Derive Dependent Inversion iinv with (forall n i I (A : PTm n) PA, InterpExt i I A PA) Sort Prop.
+
+Lemma InterpExt_cumulative n i j I (A : PTm n) PA :
+  i <= j ->
+   ⟦ A ⟧ i ;; I ↘ PA ->
+   ⟦ A ⟧ j ;; I ↘ PA.
+Proof.
+  move => h h0.
+  elim : A PA /h0;
+    hauto l:on ctrs:InterpExt solve+:(by lia).
+Qed.
+
+Lemma InterpUniv_cumulative n i (A : PTm n) 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) :=
+  (forall a, P a -> SN a) /\
+    (forall a, SNe a -> P a).
+
+Lemma adequacy_ext i n I A PA
+  (hI0 : forall j,  j < i -> forall a b, (TRedSN a b) ->  I j b -> I j a)
+  (hI : forall j, j < i -> CR (I j))
+  (h :  ⟦ A : PTm n ⟧ i ;; I ↘ PA) :
+  CR PA /\ SN A.
+Proof.
+  elim : A PA / h.
+  - hauto lq:on ctrs:SN unfold:CR.
+  - move => p A B PA PF hPA [ihA0 ihA1] hTot hRes ihPF.
+    have x : fin n by admit.
+    set Bot := VarPTm x.
+    have hb : PA Bot by hauto q:on ctrs:SNe.
+    rewrite /CR.
+    repeat split.
+    + case : p =>//=.
+      * rewrite /ProdSpace.
+        qauto use:SN_AppInv unfold:CR.
+      * rewrite /SumSpace => a []; first by apply N_SNe.
+        move => [q0][q1]*.
+        have : SN q0 /\ SN q1 by hauto q:on.

From 38a0416b2c35eb47e87d7626ece530d4e1bbead3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 4 Feb 2025 22:14:27 -0500
Subject: [PATCH 022/210] Add a constant to avoid kripke LR

---
 syntax.sig                   |  3 ++-
 theories/Autosubst2/syntax.v | 24 +++++++++++++++++++++---
 theories/fp_red.v            | 26 +++++++++++++++++++++++---
 3 files changed, 46 insertions(+), 7 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index e17d7ea..6b7e4df 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -15,4 +15,5 @@ PApp : PTm -> PTm -> PTm
 PPair : PTm -> PTm -> PTm
 PProj : PTag -> PTm -> PTm
 PBind : BTag -> PTm -> (bind PTm in PTm) -> PTm
-PUniv : nat -> PTm
\ No newline at end of file
+PUniv : nat -> PTm
+PBot : PTm
\ No newline at end of file
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index fbdf45d..ff9ec18 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -40,7 +40,8 @@ Inductive PTm (n_PTm : nat) : Type :=
   | 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.
+  | PUniv : nat -> PTm n_PTm
+  | PBot : PTm n_PTm.
 
 Lemma congr_PAbs {m_PTm : nat} {s0 : PTm (S m_PTm)} {t0 : PTm (S m_PTm)}
   (H0 : s0 = t0) : PAbs m_PTm s0 = PAbs m_PTm t0.
@@ -89,6 +90,11 @@ Proof.
 exact (eq_trans eq_refl (ap (fun x => PUniv m_PTm x) H0)).
 Qed.
 
+Lemma congr_PBot {m_PTm : nat} : PBot m_PTm = PBot m_PTm.
+Proof.
+exact (eq_refl).
+Qed.
+
 Lemma upRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n) :
   fin (S m) -> fin (S n).
 Proof.
@@ -112,6 +118,7 @@ Fixpoint ren_PTm {m_PTm : nat} {n_PTm : nat}
   | 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
   end.
 
 Lemma up_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm) :
@@ -142,6 +149,7 @@ Fixpoint subst_PTm {m_PTm : nat} {n_PTm : nat}
       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
   end.
 
 Lemma upId_PTm_PTm {m_PTm : nat} (sigma : fin m_PTm -> PTm m_PTm)
@@ -184,6 +192,7 @@ Fixpoint idSubst_PTm {m_PTm : nat} (sigma_PTm : fin m_PTm -> PTm m_PTm)
       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
   end.
 
 Lemma upExtRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n)
@@ -229,6 +238,7 @@ ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
         (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
   end.
 
 Lemma upExt_PTm_PTm {m : nat} {n_PTm : nat} (sigma : fin m -> PTm n_PTm)
@@ -275,6 +285,7 @@ subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
         (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
   end.
 
 Lemma up_ren_ren_PTm_PTm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
@@ -322,6 +333,7 @@ Fixpoint compRenRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
         (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
   end.
 
 Lemma up_ren_subst_PTm_PTm {k : nat} {l : nat} {m_PTm : nat}
@@ -378,6 +390,7 @@ Fixpoint compRenSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
         (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
   end.
 
 Lemma up_subst_ren_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
@@ -454,6 +467,7 @@ ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
         (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
   end.
 
 Lemma up_subst_subst_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
@@ -532,6 +546,7 @@ subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
         (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
   end.
 
 Lemma renRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
@@ -649,6 +664,7 @@ Fixpoint rinst_inst_PTm {m_PTm : nat} {n_PTm : nat}
         (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
   end.
 
 Lemma rinstInst'_PTm {m_PTm : nat} {n_PTm : nat}
@@ -855,6 +871,8 @@ Core.
 
 Arguments VarPTm {n_PTm}.
 
+Arguments PBot {n_PTm}.
+
 Arguments PUniv {n_PTm}.
 
 Arguments PBind {n_PTm}.
@@ -867,9 +885,9 @@ Arguments PApp {n_PTm}.
 
 Arguments PAbs {n_PTm}.
 
-#[global]Hint Opaque subst_PTm: rewrite.
+#[global] Hint Opaque subst_PTm: rewrite.
 
-#[global]Hint Opaque ren_PTm: rewrite.
+#[global] Hint Opaque ren_PTm: rewrite.
 
 End Extra.
 
diff --git a/theories/fp_red.v b/theories/fp_red.v
index ec43c8f..ac5ec3d 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -53,7 +53,9 @@ Module EPar.
   | BindCong p A0 A1 B0 B1 :
     R A0 A1 ->
     R B0 B1 ->
-    R (PBind p A0 B0) (PBind p A1 B1).
+    R (PBind p A0 B0) (PBind p A1 B1)
+  | BotCong :
+    R PBot PBot.
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -117,6 +119,8 @@ Inductive SNe {n} : PTm n -> Prop :=
 | N_Proj p a :
   SNe a ->
   SNe (PProj p a)
+| N_Bot :
+  SNe PBot
 with SN  {n} : PTm n -> Prop :=
 | N_Pair a b :
   SN a ->
@@ -270,6 +274,7 @@ Fixpoint ne {n} (a : PTm n) :=
   | PProj _ a => ne a
   | PUniv _ => false
   | PBind _ _ _ => false
+  | PBot => true
   end
 with nf {n} (a : PTm n) :=
   match a with
@@ -280,6 +285,7 @@ with nf {n} (a : PTm n) :=
   | PProj _ a => ne a
   | PUniv _ => true
   | PBind _ A B => nf A && nf B
+  | PBot => true
 end.
 
 Lemma ne_nf n a : @ne n a -> nf a.
@@ -427,6 +433,7 @@ Proof.
   - sauto lq:on.
   - sauto lq:on.
   - sauto lq:on.
+  - sauto lq:on.
   - move => a b ha iha hb ihb b0.
     inversion 1; subst.
     + have /iha : (EPar.R (PProj PL a0) (PProj PL b0)) by sauto lq:on.
@@ -595,7 +602,9 @@ Module RPar.
   | BindCong p A0 A1 B0 B1 :
     R A0 A1 ->
     R B0 B1 ->
-    R (PBind p A0 B0) (PBind p A1 B1).
+    R (PBind p A0 B0) (PBind p A1 B1)
+  | BotCong :
+    R PBot PBot.
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -690,6 +699,7 @@ 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
     end.
 
   Lemma triangle n (a b : PTm n) :
@@ -720,6 +730,7 @@ Module RPar.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
+    - hauto lq:on ctrs:R inv:R.
   Qed.
 
   Lemma diamond n (a b c : PTm n) :
@@ -736,6 +747,7 @@ Proof.
   - 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 ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN.
@@ -874,6 +886,8 @@ 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
   with R_elim {n} : PTm n -> PTm n -> Prop :=
   | NAbsCong a0 a1 :
     R_nonelim a0 a1 ->
@@ -896,7 +910,9 @@ Module NeEPar.
   | NBindCong p A0 A1 B0 B1 :
     R_nonelim A0 A1 ->
     R_nonelim B0 B1 ->
-    R_elim (PBind p A0 B0) (PBind p A1 B1).
+    R_elim (PBind p A0 B0) (PBind p A1 B1)
+  | NBotCong :
+    R_elim PBot PBot.
 
   Scheme epar_elim_ind := Induction for R_elim Sort Prop
       with epar_nonelim_ind := Induction for R_nonelim Sort Prop.
@@ -1165,6 +1181,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto lq:on ctrs:rtc, NeEPar.R_nonelim.
     - 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.
   Qed.
 
 
@@ -1272,6 +1289,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - hauto lq:on inv:RRed.R.
     - 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.
   Qed.
 
   Lemma η_postponement_star n a b c :
@@ -1762,6 +1780,7 @@ 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 q:on use:LoReds.PairCong solve+:triv.
     - hauto q:on use:LoReds.AbsCong solve+:triv.
     - sfirstorder use:ne_nf.
@@ -1797,6 +1816,7 @@ Fixpoint size_PTm {n} (a : PTm n) :=
   | PPair a b => 3 + Nat.add (size_PTm a) (size_PTm b)
   | PUniv _ => 3
   | PBind p A B => 3 + Nat.add (size_PTm A) (size_PTm B)
+  | PBot => 1
   end.
 
 Lemma size_PTm_ren n m (ξ : fin n -> fin m) a : size_PTm (ren_PTm ξ a) = size_PTm a.

From 2393cc5103db68a39522f20d8678b551b051b303 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 14:04:44 -0500
Subject: [PATCH 023/210] Finish adequacy ext

---
 theories/logrel.v | 28 +++++++++++++++++++++++++---
 1 file changed, 25 insertions(+), 3 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 1e638e6..b070751 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -118,6 +118,14 @@ Definition CR {n} (P : PTm n -> Prop) :=
   (forall a, P a -> SN a) /\
     (forall a, SNe a -> P a).
 
+Lemma N_Exps n (a b : PTm n) :
+  rtc TRedSN a b ->
+  SN b ->
+  SN a.
+Proof.
+  induction 1; eauto using N_Exp.
+Qed.
+
 Lemma adequacy_ext i n I A PA
   (hI0 : forall j,  j < i -> forall a b, (TRedSN a b) ->  I j b -> I j a)
   (hI : forall j, j < i -> CR (I j))
@@ -127,9 +135,8 @@ Proof.
   elim : A PA / h.
   - hauto lq:on ctrs:SN unfold:CR.
   - move => p A B PA PF hPA [ihA0 ihA1] hTot hRes ihPF.
-    have x : fin n by admit.
-    set Bot := VarPTm x.
-    have hb : PA Bot by hauto q:on ctrs:SNe.
+    have hb : PA PBot by hauto q:on ctrs:SNe.
+    have hb' : SN PBot by hauto q:on ctrs:SN, SNe.
     rewrite /CR.
     repeat split.
     + case : p =>//=.
@@ -138,3 +145,18 @@ Proof.
       * rewrite /SumSpace => a []; first by apply N_SNe.
         move => [q0][q1]*.
         have : SN q0 /\ SN q1 by hauto q:on.
+        hauto lq:on use:N_Pair,N_Exps.
+    + move => a ha.
+      case : p=>/=.
+      * rewrite /ProdSpace => a0 *.
+        suff : SNe (PApp a a0) by sfirstorder.
+        hauto q:on use:N_App.
+      * rewrite /SumSpace. tauto.
+    + apply N_Bind=>//=.
+      have : SN (PApp (PAbs B) PBot).
+      apply : N_Exp; eauto using N_β.
+      hauto lq:on.
+      qauto l:on use:SN_AppInv, SN_NoForbid.P_AbsInv.
+  - sfirstorder.
+  - hauto l:on ctrs:SN unfold:CR.
+Qed.

From 7f29fe03473919b7219a8132fc6002eb65233527 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 14:44:26 -0500
Subject: [PATCH 024/210] Add induction principle for InterpUniv

---
 theories/logrel.v | 101 ++++++++++++++++++++++++++++++++++++++++++----
 1 file changed, 93 insertions(+), 8 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index b070751..4c01f0d 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -13,7 +13,7 @@ Definition ProdSpace {n} (PA : PTm n -> Prop)
 
 Definition SumSpace {n} (PA : PTm n -> Prop)
   (PF : PTm n -> (PTm n -> Prop) -> Prop) t : Prop :=
-  SNe t \/ exists a b, rtc TRedSN t (PPair a b) /\ PA a /\ (forall PB, PF a PB -> PB b).
+  (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.
 
@@ -22,7 +22,7 @@ Reserved Notation "⟦ A ⟧ i ;; I ↘ S" (at level 70).
 Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop) -> Prop :=
 | InterpExt_Ne A :
   SNe A ->
-  ⟦ A ⟧ i ;; I ↘ SNe
+  ⟦ A ⟧ i ;; I ↘ (fun a => exists v, rtc TRedSN a v /\ SNe v)
 | InterpExt_Bind p A B PA PF :
   ⟦ A ⟧ i ;; I ↘ PA ->
   (forall a, PA a -> exists PB, PF a PB) ->
@@ -95,6 +95,25 @@ Qed.
 
 #[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv.
 
+Lemma InterpUniv_ind
+     : forall (n i : nat) (P : PTm n -> (PTm n -> Prop) -> Prop),
+       (forall A : PTm n, SNe A -> P A (fun a : PTm n => exists v : PTm n, rtc TRedSN a v /\ SNe v)) ->
+       (forall (p : BTag) (A : PTm n) (B : PTm (S n)) (PA : PTm n -> Prop)
+          (PF : PTm n -> (PTm n -> Prop) -> Prop),
+        ⟦ A ⟧ i ↘ PA ->
+        P 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 (subst_PTm (scons a VarPTm) B) PB) ->
+        P (PBind p A B) (BindSpace p PA PF)) ->
+       (forall j : nat, j < i -> P (PUniv j) (fun A => exists PA, ⟦ A ⟧ j ↘ PA)) ->
+       (forall (A A0 : PTm n) (PA : PTm n -> Prop), TRedSN A A0 -> ⟦ A0 ⟧ i ↘ PA -> P A0 PA -> P A PA) ->
+       forall (p : PTm n) (P0 : PTm n -> Prop), ⟦ p ⟧ i ↘ P0 -> P p P0.
+Proof.
+  elim /Wf_nat.lt_wf_ind => n ih i . simp InterpUniv.
+  apply InterpExt_ind.
+Qed.
+
 Derive Dependent Inversion iinv with (forall n i I (A : PTm n) PA, InterpExt i I A PA) Sort Prop.
 
 Lemma InterpExt_cumulative n i j I (A : PTm n) PA :
@@ -133,7 +152,7 @@ Lemma adequacy_ext i n I A PA
   CR PA /\ SN A.
 Proof.
   elim : A PA / h.
-  - hauto lq:on ctrs:SN unfold:CR.
+  - hauto l:on use:N_Exps ctrs:SN,SNe.
   - move => p A B PA PF hPA [ihA0 ihA1] hTot hRes ihPF.
     have hb : PA PBot by hauto q:on ctrs:SNe.
     have hb' : SN PBot by hauto q:on ctrs:SN, SNe.
@@ -142,16 +161,13 @@ Proof.
     + case : p =>//=.
       * rewrite /ProdSpace.
         qauto use:SN_AppInv unfold:CR.
-      * rewrite /SumSpace => a []; first by apply N_SNe.
-        move => [q0][q1]*.
-        have : SN q0 /\ SN q1 by hauto q:on.
-        hauto lq:on use:N_Pair,N_Exps.
+      * hauto q:on unfold:SumSpace use:N_SNe, N_Pair,N_Exps.
     + move => a ha.
       case : p=>/=.
       * rewrite /ProdSpace => a0 *.
         suff : SNe (PApp a a0) by sfirstorder.
         hauto q:on use:N_App.
-      * rewrite /SumSpace. tauto.
+      * sfirstorder.
     + apply N_Bind=>//=.
       have : SN (PApp (PAbs B) PBot).
       apply : N_Exp; eauto using N_β.
@@ -160,3 +176,72 @@ Proof.
   - sfirstorder.
   - hauto l:on ctrs:SN unfold:CR.
 Qed.
+
+Lemma InterpExt_Steps i n I A A0 PA :
+  rtc TRedSN A A0 ->
+  ⟦ A0 : PTm n ⟧ i ;; I ↘ PA ->
+  ⟦ A ⟧ i ;; I ↘ PA.
+Proof. induction 1; eauto using InterpExt_Step. Qed.
+
+Lemma InterpUniv_Steps i n A A0 PA :
+  rtc TRedSN A A0 ->
+  ⟦ A0 : PTm n ⟧ i ↘ PA ->
+  ⟦ A ⟧ i ↘ PA.
+Proof. hauto l:on use:InterpExt_Steps rew:db:InterpUniv. Qed.
+
+Lemma adequacy i n A PA
+  (h :  ⟦ A : PTm n ⟧ i ↘ PA) :
+  CR PA /\ SN A.
+Proof.
+  move : i A PA h.
+  elim /Wf_nat.lt_wf_ind => i ih A PA.
+  simp InterpUniv.
+  apply adequacy_ext.
+  hauto lq:on ctrs:rtc use:InterpUniv_Steps.
+  hauto l:on use:InterpExt_Ne rew:db:InterpUniv.
+Qed.
+
+Lemma InterpExt_back_clos n i I (A : PTm n) PA
+  (hI1 : forall j, j < i -> CR (I j))
+  (hI : forall j,  j < i -> forall a b, TRedSN a b ->  I j b -> I j a) :
+    ⟦ A ⟧ i ;; I ↘ PA ->
+    forall a b, TRedSN a b ->
+           PA b -> PA a.
+Proof.
+  move => h.
+  elim : A PA /h; eauto.
+  hauto q:on ctrs:rtc.
+
+  move => p A B PA PF hPA ihPA hTot hRes ihPF a b hr.
+  case : p => //=.
+  - rewrite /ProdSpace.
+    move => hba a0 PB ha hPB.
+    suff : TRedSN  (PApp a a0) (PApp b a0) by hauto lq:on.
+    apply N_AppL => //=.
+    hauto q:on use:adequacy_ext.
+  - hauto lq:on ctrs:rtc unfold:SumSpace.
+Qed.
+
+Lemma InterpUniv_back_clos n i (A : PTm n) PA :
+    ⟦ A ⟧ i ↘ PA ->
+    forall a b, TRedSN a b ->
+           PA b -> PA a.
+Proof.
+  simp InterpUniv. apply InterpExt_back_clos;
+  hauto l:on use:adequacy unfold:CR ctrs:InterpExt rew:db:InterpUniv.
+Qed.
+
+Lemma InterpUniv_back_closs n i (A : PTm n) PA :
+    ⟦ A ⟧ i ↘ PA ->
+    forall a b, rtc TRedSN a b ->
+           PA b -> PA a.
+Proof.
+  induction 2; hauto lq:on ctrs:rtc use:InterpUniv_back_clos.
+Qed.
+
+Lemma InterpExt_Join n i I (A B : PTm n) PA PB :
+  ⟦ A ⟧ i ;; I ↘ PA ->
+  ⟦ B ⟧ i ;; I ↘ PB ->
+  DJoin.R A B ->
+  PA = PB.
+Proof.

From 0e254c5ac36bcf4c8ca1acac84d0175cf2853aca Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 15:47:51 -0500
Subject: [PATCH 025/210] Start the proof that joinability preserves meaning

---
 theories/fp_red.v |   5 ++
 theories/logrel.v | 130 ++++++++++++++++++++++++----------------------
 2 files changed, 73 insertions(+), 62 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index ac5ec3d..5048bab 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -565,6 +565,11 @@ Module RRed.
     R a b -> False.
   Proof. move/[swap]. induction 1; hauto qb:on inv:PTm. Qed.
 
+  Lemma FromRedSN n (a b : PTm n) :
+    TRedSN a b ->
+    RRed.R a b.
+  Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
+
 End RRed.
 
 Module RPar.
diff --git a/theories/logrel.v b/theories/logrel.v
index 4c01f0d..76e7746 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -96,22 +96,24 @@ Qed.
 #[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv.
 
 Lemma InterpUniv_ind
-     : forall (n i : nat) (P : PTm n -> (PTm n -> Prop) -> Prop),
-       (forall A : PTm n, SNe A -> P A (fun a : PTm n => exists v : PTm n, rtc TRedSN a v /\ SNe v)) ->
-       (forall (p : BTag) (A : PTm n) (B : PTm (S n)) (PA : PTm n -> Prop)
+  : 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),
         ⟦ A ⟧ i ↘ PA ->
-        P A 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 (subst_PTm (scons a VarPTm) B) PB) ->
-        P (PBind p A B) (BindSpace p PA PF)) ->
-       (forall j : nat, j < i -> P (PUniv j) (fun A => exists PA, ⟦ A ⟧ j ↘ PA)) ->
-       (forall (A A0 : PTm n) (PA : PTm n -> Prop), TRedSN A A0 -> ⟦ A0 ⟧ i ↘ PA -> P A0 PA -> P A PA) ->
-       forall (p : PTm n) (P0 : PTm n -> Prop), ⟦ p ⟧ i ↘ P0 -> P p P0.
+        (forall (a : PTm n) (PB : PTm n -> 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 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.
 Proof.
-  elim /Wf_nat.lt_wf_ind => n ih i . simp InterpUniv.
-  apply InterpExt_ind.
+  move => n P  hSN hBind 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.
@@ -145,15 +147,13 @@ Proof.
   induction 1; eauto using N_Exp.
 Qed.
 
-Lemma adequacy_ext i n I A PA
-  (hI0 : forall j,  j < i -> forall a b, (TRedSN a b) ->  I j b -> I j a)
-  (hI : forall j, j < i -> CR (I j))
-  (h :  ⟦ A : PTm n ⟧ i ;; I ↘ PA) :
+Lemma adequacy : forall i n A PA,
+   ⟦ A : PTm n ⟧ i ↘ PA ->
   CR PA /\ SN A.
 Proof.
-  elim : A PA / h.
+  move => + n. apply : InterpUniv_ind.
   - hauto l:on use:N_Exps ctrs:SN,SNe.
-  - move => p A B PA PF hPA [ihA0 ihA1] hTot hRes ihPF.
+  - move => i p A B PA PF hPA [ihA0 ihA1] hTot hRes ihPF.
     have hb : PA PBot by hauto q:on ctrs:SNe.
     have hb' : SN PBot by hauto q:on ctrs:SN, SNe.
     rewrite /CR.
@@ -173,62 +173,38 @@ Proof.
       apply : N_Exp; eauto using N_β.
       hauto lq:on.
       qauto l:on use:SN_AppInv, SN_NoForbid.P_AbsInv.
-  - sfirstorder.
+  - hauto l:on ctrs:InterpExt rew:db:InterpUniv.
   - hauto l:on ctrs:SN unfold:CR.
 Qed.
 
-Lemma InterpExt_Steps i n I A A0 PA :
-  rtc TRedSN A A0 ->
-  ⟦ A0 : PTm n ⟧ i ;; I ↘ PA ->
-  ⟦ A ⟧ i ;; I ↘ PA.
-Proof. induction 1; eauto using InterpExt_Step. Qed.
+Lemma InterpUniv_Step i n A A0 PA :
+  TRedSN A A0 ->
+  ⟦ A0 : PTm n ⟧ i ↘ PA ->
+  ⟦ A ⟧ i ↘ PA.
+Proof. simp InterpUniv. apply InterpExt_Step. Qed.
 
 Lemma InterpUniv_Steps i n A A0 PA :
   rtc TRedSN A A0 ->
   ⟦ A0 : PTm n ⟧ i ↘ PA ->
   ⟦ A ⟧ i ↘ PA.
-Proof. hauto l:on use:InterpExt_Steps rew:db:InterpUniv. Qed.
-
-Lemma adequacy i n A PA
-  (h :  ⟦ A : PTm n ⟧ i ↘ PA) :
-  CR PA /\ SN A.
-Proof.
-  move : i A PA h.
-  elim /Wf_nat.lt_wf_ind => i ih A PA.
-  simp InterpUniv.
-  apply adequacy_ext.
-  hauto lq:on ctrs:rtc use:InterpUniv_Steps.
-  hauto l:on use:InterpExt_Ne rew:db:InterpUniv.
-Qed.
-
-Lemma InterpExt_back_clos n i I (A : PTm n) PA
-  (hI1 : forall j, j < i -> CR (I j))
-  (hI : forall j,  j < i -> forall a b, TRedSN a b ->  I j b -> I j a) :
-    ⟦ A ⟧ i ;; I ↘ PA ->
-    forall a b, TRedSN a b ->
-           PA b -> PA a.
-Proof.
-  move => h.
-  elim : A PA /h; eauto.
-  hauto q:on ctrs:rtc.
-
-  move => p A B PA PF hPA ihPA hTot hRes ihPF a b hr.
-  case : p => //=.
-  - rewrite /ProdSpace.
-    move => hba a0 PB ha hPB.
-    suff : TRedSN  (PApp a a0) (PApp b a0) by hauto lq:on.
-    apply N_AppL => //=.
-    hauto q:on use:adequacy_ext.
-  - hauto lq:on ctrs:rtc unfold:SumSpace.
-Qed.
+Proof. induction 1; hauto l:on use:InterpUniv_Step. Qed.
 
 Lemma InterpUniv_back_clos n i (A : PTm n) PA :
     ⟦ A ⟧ i ↘ PA ->
     forall a b, TRedSN a b ->
            PA b -> PA a.
 Proof.
-  simp InterpUniv. apply InterpExt_back_clos;
-  hauto l:on use:adequacy unfold:CR ctrs:InterpExt rew:db:InterpUniv.
+  move : i A PA . apply : InterpUniv_ind; eauto.
+  - hauto q:on ctrs:rtc.
+  - move => i p A B PA PF hPA ihPA hTot hRes ihPF a b hr.
+    case : p => //=.
+    + rewrite /ProdSpace.
+    move => hba a0 PB ha hPB.
+    suff : TRedSN  (PApp a a0) (PApp b a0) by hauto lq:on.
+    apply N_AppL => //=.
+    hauto q:on use:adequacy.
+    + hauto lq:on ctrs:rtc unfold:SumSpace.
+  - hauto l:on use:InterpUniv_Step.
 Qed.
 
 Lemma InterpUniv_back_closs n i (A : PTm n) PA :
@@ -239,9 +215,39 @@ Proof.
   induction 2; hauto lq:on ctrs:rtc use:InterpUniv_back_clos.
 Qed.
 
-Lemma InterpExt_Join n i I (A B : PTm n) PA PB :
-  ⟦ A ⟧ i ;; I ↘ PA ->
-  ⟦ B ⟧ i ;; I ↘ PB ->
+Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
+
+Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
+
+Lemma InterpUniv_case n i (A : PTm n) PA :
+  ⟦ A ⟧ i ↘ PA ->
+  exists H, rtc TRedSN A H /\ (SNe H \/ isbind H \/ isuniv H).
+Proof.
+  move : i A PA. apply InterpUniv_ind => //=; hauto ctrs:rtc l:on.
+Qed.
+
+Lemma redsn_preservation_mutual n :
+  (forall (a : PTm n) (s : SNe a), forall b, TRedSN a b -> SNe b) /\
+    (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).
+Proof.
+  move : n. 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.
+Proof. induction 2; sfirstorder use:redsn_preservation_mutual ctrs:rtc. Qed.
+
+Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
+  ⟦ A ⟧ i ↘ PA ->
+  ⟦ B ⟧ i ↘ PB ->
   DJoin.R A B ->
   PA = PB.
 Proof.
+  move => hA.
+  move : i A PA hA B PB.
+  apply : InterpUniv_ind.
+  - move => i A hA B PB hPB hAB.
+    have [*] : SN B /\ SN A by hauto l:on use:adequacy.
+    move /InterpUniv_case : hPB.
+    move => [H [h ?]].
+    (* have ? : SN H by sfirstorder use:redsns_preservation. *)

From e444c8408f256a185391ba9f5552c25676abb604 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 16:52:25 -0500
Subject: [PATCH 026/210] Show that sne and bind are not joinable

---
 theories/fp_red.v | 56 ++++++++++++++++++++++++++++++++++++++++++++++-
 theories/logrel.v | 12 +++++-----
 2 files changed, 62 insertions(+), 6 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5048bab..4fd2ed4 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -170,6 +170,9 @@ Definition ishf {n} (a : PTm n) :=
   | PBind _ _ _ => true
   | _ => false
   end.
+Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
+
+Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
 
 Definition ispair {n} (a : PTm n) :=
   match a with
@@ -195,7 +198,6 @@ Definition ispair_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
   ispair (ren_PTm ξ a) = ispair a.
 Proof. case : a => //=. Qed.
 
-
 Lemma PProj_imp n p a :
   @ishf n a ->
   ~~ ispair a ->
@@ -848,6 +850,11 @@ Module RReds.
     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) :
+    rtc TRedSN a b ->
+    rtc RRed.R a b.
+  Proof. induction 1; hauto lq:on ctrs:rtc use:RRed.FromRedSN. Qed.
+
 End RReds.
 
 
@@ -1534,6 +1541,11 @@ Module EReds.
       apply ERed.PairEta.
   Qed.
 
+  Lemma FromEPars n (a b : PTm n) :
+    rtc EPar.R a b ->
+    rtc ERed.R a b.
+  Proof. induction 1; hauto l:on use:FromEPar, @relations.rtc_transitive. Qed.
+
 End EReds.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -1606,6 +1618,20 @@ Module RERed.
     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) :
+    R a b -> isbind a -> isbind b.
+  Proof. hauto q:on inv:R. Qed.
+
+  Lemma univ_preservation n (a b : PTm n) :
+    R a b -> isuniv a -> isuniv b.
+  Proof. hauto q:on inv:R. Qed.
+
+  Lemma sne_preservation n (a b : PTm n) :
+    R a b -> SNe a -> SNe b.
+  Proof.
+    hauto q:on use:ToBetaEtaPar, RPar.FromRRed use:red_sn_preservation, epar_sn_preservation.
+  Qed.
+
 End RERed.
 
 Module REReds.
@@ -1660,6 +1686,18 @@ Module REReds.
     rtc RERed.R (PBind p A0 B0) (PBind p A1 B1).
   Proof. solve_s. Qed.
 
+  Lemma bind_preservation n (a b : PTm n) :
+    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) :
+    rtc RERed.R a b -> isuniv a -> isuniv b.
+  Proof. induction 1; qauto l:on ctrs:rtc use:RERed.univ_preservation. Qed.
+
+  Lemma sne_preservation n (a b : PTm n) :
+    rtc RERed.R a b -> SNe a -> SNe b.
+  Proof. induction 1; qauto l:on ctrs:rtc use:RERed.sne_preservation. Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2051,4 +2089,20 @@ Module DJoin.
     R (PProj p a0) (PProj p a1).
   Proof. hauto q:on use:REReds.ProjCong. Qed.
 
+  Lemma FromRedSNs n (a b : PTm n) :
+    rtc TRedSN a b ->
+    R a b.
+  Proof.
+    move /RReds.FromRedSNs /REReds.FromRReds.
+    sfirstorder use:@rtc_refl unfold:R.
+  Qed.
+
+  Lemma sne_bind_imp n (a b : PTm n) :
+    R a b -> SNe a -> isbind b -> False.
+  Proof.
+    move => [c [? ?]] *.
+    have : SNe c /\ isbind c by sfirstorder use:REReds.sne_preservation, REReds.bind_preservation.
+    qauto l:on inv:SNe.
+  Qed.
+
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index 76e7746..424351b 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -215,9 +215,6 @@ Proof.
   induction 2; hauto lq:on ctrs:rtc use:InterpUniv_back_clos.
 Qed.
 
-Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
-
-Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
 
 Lemma InterpUniv_case n i (A : PTm n) PA :
   ⟦ A ⟧ i ↘ PA ->
@@ -237,6 +234,11 @@ Qed.
 Lemma redsns_preservation : forall n a b, @SN n a -> rtc TRedSN a b -> SN b.
 Proof. induction 2; sfirstorder use:redsn_preservation_mutual ctrs:rtc. Qed.
 
+Lemma sne_bind_noconf n (a b : PTm n) :
+  SNe a -> isbind b -> DJoin.R a b -> False.
+Proof.
+
+
 Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
@@ -249,5 +251,5 @@ Proof.
   - move => i A hA B PB hPB hAB.
     have [*] : SN B /\ SN A by hauto l:on use:adequacy.
     move /InterpUniv_case : hPB.
-    move => [H [h ?]].
-    (* have ? : SN H by sfirstorder use:redsns_preservation. *)
+    move => [H [/DJoin.FromRedSNs h ?]].
+    have ? : DJoin.R A H by eauto using DJoin.transitive.

From af224831e408333187b374dc7ea25c79cea383b0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 5 Feb 2025 18:56:47 -0500
Subject: [PATCH 027/210] Finish the injectivity of bind and noconfusion

---
 theories/fp_red.v | 36 +++++++++++++++++++++++++++++++++++-
 theories/logrel.v |  1 +
 2 files changed, 36 insertions(+), 1 deletion(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 4fd2ed4..5178ab5 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1698,6 +1698,15 @@ Module REReds.
     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 ->
+    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.
+    move : p A B E.
+    elim : u C / hu; sauto lq:on rew:off.
+  Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2097,7 +2106,7 @@ Module DJoin.
     sfirstorder use:@rtc_refl unfold:R.
   Qed.
 
-  Lemma sne_bind_imp n (a b : PTm n) :
+  Lemma sne_bind_noconf n (a b : PTm n) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
     move => [c [? ?]] *.
@@ -2105,4 +2114,29 @@ Module DJoin.
     qauto l:on inv:SNe.
   Qed.
 
+  Lemma sne_univ_noconf n (a b : PTm n) :
+    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) :
+    R a b -> isbind a -> isuniv b -> False.
+  Proof.
+    move => [c [h0 h1]] h2 h3.
+    have {h0 h1 h2 h3} : isbind c /\ isuniv c by
+      hauto l:on use:REReds.bind_preservation,
+          REReds.univ_preservation.
+    case : c => //=; itauto.
+  Qed.
+
+  Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
+    DJoin.R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
+    p0 = p1 /\ DJoin.R A0 A1 /\ DJoin.R B0 B1.
+  Proof.
+    rewrite /R.
+    hauto lq:on rew:off use:REReds.bind_inv.
+  Qed.
+
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index 424351b..effd176 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -239,6 +239,7 @@ Lemma sne_bind_noconf n (a b : PTm n) :
 Proof.
 
 
+
 Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->

From 7cc6435ea3e039b7db59979d63fdc578217c6896 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 5 Feb 2025 20:06:03 -0500
Subject: [PATCH 028/210] Finish most of InterpUniv join

---
 theories/logrel.v | 118 ++++++++++++++++++++++++++++++++++++++++------
 1 file changed, 104 insertions(+), 14 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index effd176..3686f81 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -6,6 +6,7 @@ Require Import ssreflect ssrbool.
 Require Import Logic.PropExtensionality (propositional_extensionality).
 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 :=
@@ -118,6 +119,34 @@ Qed.
 
 Derive Dependent Inversion iinv with (forall n i I (A : PTm n) PA, InterpExt i I A PA) Sort Prop.
 
+Lemma InterpUniv_Ne n i (A : PTm n) :
+  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 ->
+  (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).
+Proof.
+  simp InterpUniv. simpl.
+  apply InterpExt_Univ'. by simp InterpUniv.
+Qed.
+
+Lemma InterpUniv_Step i n A A0 PA :
+  TRedSN A A0 ->
+  ⟦ A0 : PTm n ⟧ i ↘ PA ->
+  ⟦ A ⟧ i ↘ PA.
+Proof. simp InterpUniv. apply InterpExt_Step. Qed.
+
+
+#[export]Hint Resolve InterpUniv_Bind InterpUniv_Step InterpUniv_Ne InterpUniv_Univ : InterpUniv.
+
 Lemma InterpExt_cumulative n i j I (A : PTm n) PA :
   i <= j ->
    ⟦ A ⟧ i ;; I ↘ PA ->
@@ -177,12 +206,6 @@ Proof.
   - hauto l:on ctrs:SN unfold:CR.
 Qed.
 
-Lemma InterpUniv_Step i n A A0 PA :
-  TRedSN A A0 ->
-  ⟦ A0 : PTm n ⟧ i ↘ PA ->
-  ⟦ A ⟧ i ↘ PA.
-Proof. simp InterpUniv. apply InterpExt_Step. Qed.
-
 Lemma InterpUniv_Steps i n A A0 PA :
   rtc TRedSN A A0 ->
   ⟦ A0 : PTm n ⟧ i ↘ PA ->
@@ -218,13 +241,17 @@ Qed.
 
 Lemma InterpUniv_case n i (A : PTm n) PA :
   ⟦ A ⟧ i ↘ PA ->
-  exists H, rtc TRedSN A H /\ (SNe H \/ isbind H \/ isuniv H).
+  exists H, rtc TRedSN A H /\  ⟦ H ⟧ i ↘ PA /\ (SNe H \/ isbind H \/ isuniv H).
 Proof.
-  move : i A PA. apply InterpUniv_ind => //=; hauto ctrs:rtc l:on.
+  move : i A PA. apply InterpUniv_ind => //=.
+  hauto lq:on ctrs:rtc use:InterpUniv_Ne.
+  hauto l:on use:InterpUniv_Bind.
+  hauto l:on use:InterpUniv_Univ.
+  hauto lq:on ctrs:rtc.
 Qed.
 
 Lemma redsn_preservation_mutual n :
-  (forall (a : PTm n) (s : SNe a), forall b, TRedSN a b -> SNe b) /\
+  (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).
 Proof.
@@ -234,10 +261,38 @@ Qed.
 Lemma redsns_preservation : forall n a b, @SN n a -> rtc TRedSN a b -> SN b.
 Proof. induction 2; sfirstorder use:redsn_preservation_mutual ctrs:rtc. Qed.
 
-Lemma sne_bind_noconf n (a b : PTm n) :
-  SNe a -> isbind b -> DJoin.R a b -> False.
-Proof.
+#[export]Hint Resolve DJoin.sne_bind_noconf DJoin.sne_univ_noconf DJoin.bind_univ_noconf : noconf.
 
+Lemma InterpUniv_SNe_inv n i (A : PTm n) PA :
+  SNe A ->
+  ⟦ A ⟧ i ↘ PA ->
+  PA = (fun a => exists v, rtc TRedSN a v /\ SNe v).
+Proof.
+  simp InterpUniv.
+  hauto lq:on rew:off inv:InterpExt,SNe use:redsn_preservation_mutual.
+Qed.
+
+Lemma InterpUniv_Bind_inv n i p A B S :
+  ⟦ PBind p A B ⟧ i ↘ S -> exists PA PF,
+  ⟦ A : PTm n ⟧ 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.
+Proof. simp InterpUniv.
+       inversion 1; try hauto inv:SNe q:on use:redsn_preservation_mutual.
+       rewrite -!InterpUnivN_nolt.
+       sauto lq:on.
+Qed.
+
+Lemma InterpUniv_Univ_inv n i j S :
+  ⟦ PUniv j : PTm n ⟧ i ↘ S ->
+  S = (fun A => exists PA, ⟦ A ⟧ j ↘ PA) /\ j < i.
+Proof.
+  simp InterpUniv. inversion 1;
+    try hauto inv:SNe use:redsn_preservation_mutual.
+  rewrite -!InterpUnivN_nolt. sfirstorder.
+  subst. hauto lq:on inv:TRedSN.
+Qed.
 
 
 Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
@@ -252,5 +307,40 @@ Proof.
   - move => i A hA B PB hPB hAB.
     have [*] : SN B /\ SN A by hauto l:on use:adequacy.
     move /InterpUniv_case : hPB.
-    move => [H [/DJoin.FromRedSNs h ?]].
-    have ? : DJoin.R A H by eauto using DJoin.transitive.
+    move => [H [/DJoin.FromRedSNs h [h1 h0]]].
+    have {hAB} {}h : DJoin.R A H by eauto using DJoin.transitive.
+    have {}h0 : SNe H.
+    suff : ~ isbind H /\ ~ isuniv H by itauto.
+    move : h hA. clear. hauto lq:on db:noconf.
+    hauto lq:on use:InterpUniv_SNe_inv.
+  - move => i p A B PA PF hPA ihPA hTot hRes ihPF U PU hU.
+    have hU' : SN U by hauto l:on use:adequacy.
+    move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU.
+    have {hU} {}h : DJoin.R (PBind p A B) H by eauto using DJoin.transitive.
+    have{h0} : isbind H.
+    suff : ~ SNe H /\ ~ isuniv H by itauto.
+    have : isbind (PBind p A B) by scongruence.
+    hauto l:on use: DJoin.sne_bind_noconf, DJoin.bind_univ_noconf, DJoin.symmetric.
+    case : H h1 h => //=.
+    move => p0 A0 B0 h0 /DJoin.bind_inj.
+    move => [? [hA hB]] _. subst.
+    admit.
+  - move => i j jlti ih B PB hPB.
+    have ? : SN B by hauto l:on use:adequacy.
+    move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].
+    move => hj.
+    have {hj}{}h : DJoin.R (PUniv j) H by eauto using DJoin.transitive.
+    have {h0} : isuniv H.
+    suff : ~ SNe H /\ ~ isbind H by tauto.
+    hauto l:on use: DJoin.sne_univ_noconf, DJoin.bind_univ_noconf, DJoin.symmetric.
+    case : H h1 h => //=.
+    move => j' hPB h _.
+    have {}h : j' = j by admit. subst.
+    hauto lq:on use:InterpUniv_Univ_inv.
+  - move => i A A0 PA hr hPA ihPA B PB hPB hAB.
+    have /DJoin.symmetric ? : DJoin.R A A0 by hauto lq:on rew:off ctrs:rtc use:DJoin.FromRedSNs.
+    have ? : SN A0 by hauto l:on use:adequacy.
+    have ? : SN A by eauto using N_Exp.
+    have : DJoin.R A0 B by eauto using DJoin.transitive.
+    eauto.
+Admitted.

From 1997e8bc12235daecbc913a92699af63d5c65f4e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 5 Feb 2025 20:36:39 -0500
Subject: [PATCH 029/210] Side step the need for join subst

---
 theories/fp_red.v | 12 ++++++++++++
 theories/logrel.v | 28 +++++++++++++++++++++++++++-
 2 files changed, 39 insertions(+), 1 deletion(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5178ab5..a71a0a0 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2139,4 +2139,16 @@ Module DJoin.
     hauto lq:on rew:off use:REReds.bind_inv.
   Qed.
 
+  Lemma FromRRed0 n (a b : PTm n) :
+    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) :
+    RRed.R b a -> R a b.
+  Proof.
+    hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
+  Qed.
+
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index 3686f81..bcc06be 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -324,7 +324,33 @@ Proof.
     case : H h1 h => //=.
     move => p0 A0 B0 h0 /DJoin.bind_inj.
     move => [? [hA hB]] _. subst.
-    admit.
+    move /InterpUniv_Bind_inv : h0.
+    move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst.
+    have ? : PA0 = PA by qauto l:on. subst.
+    have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'.
+    move => a PB PB' ha hPB hPB'. apply : ihPF; eauto.
+    have hj0 : DJoin.R (PAbs B) (PAbs B0) by eauto using DJoin.AbsCong.
+    have {}hj0 : DJoin.R (PApp (PAbs B) a) (PApp (PAbs B0) a) by eauto using DJoin.AppCong, DJoin.refl.
+    have ? : SN (PApp (PAbs B) a)
+      by hauto lq:on rew:off use:N_Exp, N_β, adequacy.
+    have ? : SN (PApp (PAbs B0) a)
+      by hauto lq:on rew:off use:N_Exp, N_β, adequacy.
+    have : DJoin.R (PApp (PAbs B0) a) (subst_PTm (scons a VarPTm) B0)
+      by hauto lq:on ctrs:RRed.R use:DJoin.FromRRed0.
+    have : DJoin.R (subst_PTm (scons a VarPTm) B) (PApp (PAbs B) a)
+      by hauto lq:on ctrs:RRed.R use:DJoin.FromRRed1.
+    eauto using DJoin.transitive.
+
+    move => hI.
+    case : p0 => //=.
+    + rewrite /ProdSpace.
+      extensionality b.
+      apply propositional_extensionality.
+      split.
+      move =>  hPF a PB.
+
+      move => a PB hPB.
+
   - move => i j jlti ih B PB hPB.
     have ? : SN B by hauto l:on use:adequacy.
     move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].

From 76f8c32dad6befa32f1a514f0b7b48ba8056cf0d Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 21:14:31 -0500
Subject: [PATCH 030/210] Finish the hard fun case for interpuniv_join

---
 theories/logrel.v | 47 ++++++++++++++++++++++++++++-------------------
 1 file changed, 28 insertions(+), 19 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index bcc06be..bb7571e 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -294,6 +294,27 @@ Proof.
   subst. hauto lq:on inv:TRedSN.
 Qed.
 
+Lemma bindspace_iff n p (PA : PTm n -> Prop) PF PF0 b  :
+  (forall (a : PTm n) (PB PB0 : PTm n -> Prop), PA a -> PF a PB -> PF0 a PB0 -> PB = PB0) ->
+  (forall a, PA a -> exists PB, PF a PB) ->
+  (forall a, PA a -> exists PB0, PF0 a PB0) ->
+  (BindSpace p PA PF b <-> BindSpace p PA PF0 b).
+Proof.
+  rewrite /BindSpace => h hPF hPF0.
+  case : p => /=.
+  - rewrite /ProdSpace.
+    split.
+    move => h1 a PB ha hPF'.
+    specialize hPF with (1 := ha).
+    specialize hPF0 with (1 := ha).
+    sblast.
+    move => ? a PB ha.
+    specialize hPF with (1 := ha).
+    specialize hPF0 with (1 := ha).
+    sblast.
+  - rewrite /SumSpace.
+    hauto lq:on rew:off.
+Qed.
 
 Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
@@ -331,26 +352,14 @@ Proof.
     move => a PB PB' ha hPB hPB'. apply : ihPF; eauto.
     have hj0 : DJoin.R (PAbs B) (PAbs B0) by eauto using DJoin.AbsCong.
     have {}hj0 : DJoin.R (PApp (PAbs B) a) (PApp (PAbs B0) a) by eauto using DJoin.AppCong, DJoin.refl.
-    have ? : SN (PApp (PAbs B) a)
-      by hauto lq:on rew:off use:N_Exp, N_β, adequacy.
-    have ? : SN (PApp (PAbs B0) a)
-      by hauto lq:on rew:off use:N_Exp, N_β, adequacy.
-    have : DJoin.R (PApp (PAbs B0) a) (subst_PTm (scons a VarPTm) B0)
-      by hauto lq:on ctrs:RRed.R use:DJoin.FromRRed0.
-    have : DJoin.R (subst_PTm (scons a VarPTm) B) (PApp (PAbs B) a)
-      by hauto lq:on ctrs:RRed.R use:DJoin.FromRRed1.
+    have [? ?] : SN (PApp (PAbs B) a) /\ SN (PApp (PAbs B0) a) by
+      hauto lq:on rew:off use:N_Exp, N_β, adequacy.
+    have [? ?] : DJoin.R (PApp (PAbs B0) a) (subst_PTm (scons a VarPTm) B0) /\
+           DJoin.R (subst_PTm (scons a VarPTm) B) (PApp (PAbs B) a)
+      by hauto lq:on ctrs:RRed.R use:DJoin.FromRRed0, DJoin.FromRRed1.
     eauto using DJoin.transitive.
-
-    move => hI.
-    case : p0 => //=.
-    + rewrite /ProdSpace.
-      extensionality b.
-      apply propositional_extensionality.
-      split.
-      move =>  hPF a PB.
-
-      move => a PB hPB.
-
+    move => h. extensionality b. apply propositional_extensionality.
+    hauto l:on use:bindspace_iff.
   - move => i j jlti ih B PB hPB.
     have ? : SN B by hauto l:on use:adequacy.
     move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].

From 4134fbdada8e4a85c796a952fdb90a6ef9db02d7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 21:20:37 -0500
Subject: [PATCH 031/210] Finish InterpUniv_Join

---
 theories/fp_red.v | 17 +++++++++++++++++
 theories/logrel.v |  4 ++--
 2 files changed, 19 insertions(+), 2 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index a71a0a0..9f5ea47 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1707,6 +1707,15 @@ 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  ->
+    C = PUniv i.
+  Proof.
+    move E : (PUniv i) => u hu.
+    move : i E. elim : u C / hu=>//=.
+    hauto lq:on rew:off ctrs:rtc inv:RERed.R.
+  Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2139,6 +2148,12 @@ Module DJoin.
     hauto lq:on rew:off use:REReds.bind_inv.
   Qed.
 
+  Lemma univ_inj n i j :
+    @R n (PUniv i) (PUniv j)  -> i = j.
+  Proof.
+    sauto lq:on rew:off use:REReds.univ_inv.
+  Qed.
+
   Lemma FromRRed0 n (a b : PTm n) :
     RRed.R a b -> R a b.
   Proof.
@@ -2151,4 +2166,6 @@ Module DJoin.
     hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
   Qed.
 
+
+
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index bb7571e..48cb447 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -370,7 +370,7 @@ Proof.
     hauto l:on use: DJoin.sne_univ_noconf, DJoin.bind_univ_noconf, DJoin.symmetric.
     case : H h1 h => //=.
     move => j' hPB h _.
-    have {}h : j' = j by admit. subst.
+    have {}h : j' = j by hauto lq:on use: DJoin.univ_inj. subst.
     hauto lq:on use:InterpUniv_Univ_inv.
   - move => i A A0 PA hr hPA ihPA B PB hPB hAB.
     have /DJoin.symmetric ? : DJoin.R A A0 by hauto lq:on rew:off ctrs:rtc use:DJoin.FromRedSNs.
@@ -378,4 +378,4 @@ Proof.
     have ? : SN A by eauto using N_Exp.
     have : DJoin.R A0 B by eauto using DJoin.transitive.
     eauto.
-Admitted.
+Qed.

From 55ccc2bc1dc90c02b15e0af04d38c464615cc9f7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 21:28:39 -0500
Subject: [PATCH 032/210] Prove the enhanced interpuniv bind inversion
 principle

---
 theories/logrel.v | 55 +++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 55 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index 48cb447..c6d3990 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -379,3 +379,58 @@ Proof.
     have : DJoin.R A0 B by eauto using DJoin.transitive.
     eauto.
 Qed.
+
+Lemma InterpUniv_Functional n i  (A : PTm n) PA PB :
+  ⟦ A ⟧ i ↘ PA ->
+  ⟦ A ⟧ i ↘ PB ->
+  PA = PB.
+Proof. hauto use:InterpUniv_Join, DJoin.refl. Qed.
+
+Lemma InterpUniv_Join' n i j (A B : PTm n) PA PB :
+  ⟦ A ⟧ i ↘ PA ->
+  ⟦ B ⟧ j ↘ PB ->
+  DJoin.R A B ->
+  PA = PB.
+Proof.
+  have [? ?] : i <= max i j /\ j <= max i j by lia.
+  move => hPA hPB.
+  have : ⟦ A ⟧ (max i j) ↘ PA by eauto using InterpUniv_cumulative.
+  have : ⟦ B ⟧ (max i j) ↘ PB by eauto using InterpUniv_cumulative.
+  eauto using InterpUniv_Join.
+Qed.
+
+Lemma InterpUniv_Functional' n i j A PA PB :
+  ⟦ A : PTm n ⟧ 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),
+     ⟦ 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).
+Proof.
+  move /InterpUniv_Bind_inv : h.
+  move => [PA][PF][hPA][hPF][hPF']?. subst.
+  exists PA. repeat split => //.
+  - sfirstorder.
+  - extensionality b.
+    case : p => /=.
+    + extensionality a.
+      extensionality PB.
+      extensionality ha.
+      apply propositional_extensionality.
+      split.
+      * move => h hPB.  apply h.
+        have {}/hPF := ha.
+        move => [PB0 hPB0].
+        have {}/hPF' := hPB0 => ?.
+        have : PB = PB0 by hauto l:on use:InterpUniv_Functional.
+        congruence.
+      * sfirstorder.
+    + rewrite /SumSpace. apply propositional_extensionality.
+      split; hauto q:on use:InterpUniv_Functional.
+Qed.

From d6a96430f0e1ec07a3684dde4127c0f526abd1e9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 21:44:35 -0500
Subject: [PATCH 033/210] Add semantic typing

---
 theories/logrel.v | 118 ++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 118 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index c6d3990..43270c1 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -434,3 +434,121 @@ Proof.
     + rewrite /SumSpace. apply propositional_extensionality.
       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 SemWt {n} Γ (a A : PTm n) := forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a).
+Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70).
+
+(* Semantic context wellformedness *)
+Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ ⊨ Γ i ∈ PUniv j.
+Notation "⊨ Γ" := (SemWff Γ) (at level 70).
+
+Lemma ρ_ok_bot n (Γ : fin n -> PTm n)  :
+  ρ_ok Γ (fun _ => PBot).
+Proof.
+  rewrite /ρ_ok.
+  hauto q:on use:adequacy ctrs:SNe.
+Qed.
+
+Lemma ρ_ok_cons n i (Γ : fin n -> PTm n) ρ a PA A :
+  ⟦ subst_PTm ρ A ⟧ i ↘ PA -> PA a ->
+  ρ_ok Γ ρ ->
+  ρ_ok (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ).
+Proof.
+  move => h0 h1 h2.
+  rewrite /ρ_ok.
+  move => j.
+  destruct j as [j|].
+  - move => m PA0. asimpl => ?.
+    asimpl.
+    firstorder.
+  - move => m PA0. asimpl => h3.
+    have ? : PA0 = PA by eauto using InterpUniv_Functional'.
+    by subst.
+Qed.
+
+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).
+
+Lemma ρ_ok_renaming n m (Γ : fin n -> PTm n) ρ :
+  forall (Δ : fin m -> PTm m) ξ,
+    renaming_ok Γ Δ ξ ->
+    ρ_ok Γ ρ ->
+    ρ_ok Δ (funcomp ρ ξ).
+Proof.
+  move => Δ ξ hξ hρ.
+  rewrite /ρ_ok => i m' PA.
+  rewrite /renaming_ok in hξ.
+  rewrite /ρ_ok in hρ.
+  move => h.
+  rewrite /funcomp.
+  apply hρ with (k := m').
+  move : h. rewrite -hξ.
+  by asimpl.
+Qed.
+
+Lemma renaming_SemWt {n} Γ a A :
+  Γ ⊨ a ∈ A ->
+  forall {m} Δ (ξ : fin n -> fin m),
+    renaming_ok Δ Γ ξ ->
+    Δ ⊨ ren_PTm ξ a ∈ ren_PTm ξ A.
+Proof.
+  rewrite /SemWt => h m Δ ξ hξ ρ hρ.
+  have /h hρ' : (ρ_ok Γ (funcomp ρ ξ)) by eauto using ρ_ok_renaming.
+  hauto q:on solve+:(by asimpl).
+Qed.
+
+Lemma weakening_Sem n Γ (a : PTm n) A B i
+  (h0 : Γ ⊨ B ∈ PUniv i)
+  (h1 : Γ ⊨ a ∈ A) :
+   funcomp (ren_PTm shift) (scons B Γ) ⊨ ren_PTm shift a ∈ ren_PTm shift A.
+Proof.
+  apply : renaming_SemWt; eauto.
+  hauto lq:on inv:option unfold:renaming_ok.
+Qed.
+
+Lemma SemWt_Wn n Γ (a : PTm n) 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.
+Qed.
+
+Lemma SemWt_Univ n Γ (A : PTm n) i  :
+  Γ ⊨ A ∈ PUniv i <->
+  forall ρ, ρ_ok Γ ρ -> exists S, ⟦ subst_PTm ρ A ⟧ i ↘ S.
+Proof.
+  rewrite /SemWt.
+  split.
+  - hauto lq:on rew:off use:InterpUniv_Univ_inv.
+  - move => /[swap] ρ /[apply].
+    move => [PA hPA].
+    exists (S i). eexists.
+    split.
+    + simp InterpUniv. apply InterpExt_Univ. lia.
+    + simpl. eauto.
+Qed.
+
+(* Structural laws for Semantic context wellformedness *)
+Lemma SemWff_nil : SemWff null.
+Proof. case. Qed.
+
+Lemma SemWff_cons n Γ (A : PTm n) i :
+    ⊨ Γ ->
+    Γ ⊨ A ∈ PUniv i ->
+    (* -------------- *)
+    ⊨ funcomp (ren_PTm shift) (scons 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.
+  - hauto q:on use:weakening_Sem.
+Qed.

From 64bcf8e9c14b497ad8807cf9d46faa1596570629 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 5 Feb 2025 23:56:47 -0500
Subject: [PATCH 034/210] Finish Proj case

---
 theories/fp_red.v |  41 +++++++++-
 theories/logrel.v | 185 +++++++++++++++++++++++++++++++++++++++++++++-
 2 files changed, 224 insertions(+), 2 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 9f5ea47..5af61f4 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2066,7 +2066,36 @@ Proof.
   move /REReds.FromRReds : hc0. move /REReds.FromEReds : hv'. eauto using relations.rtc_transitive.
 Qed.
 
-(* "Declarative" Joinability *)
+(* 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.
+  Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
+
+  Lemma symmetric n (a b : PTm n) : 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.
+  Proof.
+    rewrite /R.
+    move => [ab [ha +]] [bc [+ hc]].
+    move : red_confluence; repeat move/[apply].
+    move => [v [h0 h1]].
+    exists v. sfirstorder use:@relations.rtc_transitive.
+  Qed.
+
+  Lemma AbsCong n (a b : PTm (S n)) :
+    R a b ->
+    R (PAbs a) (PAbs b).
+  Proof. hauto lq:on use:RReds.AbsCong unfold:R. Qed.
+
+  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
+    R a0 a1 ->
+    R b0 b1 ->
+    R (PApp a0 b0) (PApp a1 b1).
+  Proof. hauto lq:on use:RReds.AppCong unfold:R. Qed.
+End BJoin.
+
 Module DJoin.
   Definition R {n} (a b : PTm n) := exists c, rtc RERed.R a c /\ rtc RERed.R b c.
 
@@ -2166,6 +2195,16 @@ Module DJoin.
     hauto lq:on ctrs:rtc use:RERed.FromBeta unfold:R.
   Qed.
 
+  Lemma FromRReds n (a b : PTm n) :
+    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) :
+    BJoin.R a b -> R a b.
+  Proof.
+    hauto lq:on ctrs:rtc use:REReds.FromRReds unfold:R.
+  Qed.
 
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index 43270c1..e88449b 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -509,7 +509,7 @@ Proof.
   hauto lq:on inv:option unfold:renaming_ok.
 Qed.
 
-Lemma SemWt_Wn n Γ (a : PTm n) A :
+Lemma SemWt_SN n Γ (a : PTm n) A :
   Γ ⊨ a ∈ A ->
   SN a /\ SN A.
 Proof.
@@ -552,3 +552,186 @@ Proof.
     eauto using weakening_Sem.
   - hauto q:on use:weakening_Sem.
 Qed.
+
+(* Semantic typing rules *)
+Lemma ST_Var n Γ (i : fin n) :
+  ⊨ Γ ->
+  Γ ⊨ VarPTm i ∈ Γ i.
+Proof.
+  move /(_ i) => [j /SemWt_Univ h].
+  rewrite /SemWt => ρ /[dup] hρ {}/h [S hS].
+  exists j, S.
+  asimpl. firstorder.
+Qed.
+
+Lemma InterpUniv_Bind_nopf n p i (A : PTm n) 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)).
+Proof.
+  move => h0 h1. apply InterpUniv_Bind => //=.
+Qed.
+
+
+Lemma ST_Bind n Γ i j p (A : PTm n) (B : PTm (S n)) :
+  Γ ⊨ A ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv j ->
+  Γ ⊨ PBind p A B ∈ PUniv (max i j).
+Proof.
+  move => /SemWt_Univ h0 /SemWt_Univ h1.
+  apply SemWt_Univ => ρ hρ.
+  move /h0 : (hρ){h0} => [S hS].
+  eexists => /=.
+  have ? : i <= Nat.max i j by lia.
+  apply InterpUniv_Bind_nopf; eauto.
+  - eauto using InterpUniv_cumulative.
+  - move => *. asimpl. hauto l:on use:InterpUniv_cumulative, ρ_ok_cons.
+Qed.
+
+Lemma ST_Abs n Γ (a : PTm (S n)) A B i :
+  Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊨ a ∈ B ->
+  Γ ⊨ PAbs a ∈ PBind PPi A B.
+Proof.
+  rename a into b.
+  move /SemWt_Univ => + hb ρ hρ.
+  move /(_ _ hρ) => [PPi hPPi].
+  exists i, PPi. split => //.
+  simpl in hPPi.
+  move /InterpUniv_Bind_inv_nopf : hPPi.
+  move => [PA [hPA [hTot ?]]]. subst=>/=.
+  move => a PB ha. asimpl => hPB.
+  move : ρ_ok_cons (hPA) (hρ) (ha). repeat move/[apply].
+  move /hb.
+  intros (m & PB0 & hPB0 & hPB0').
+  replace PB0 with PB in * by hauto l:on use:InterpUniv_Functional'.
+  apply : InterpUniv_back_clos; eauto.
+  apply N_β'. by asimpl.
+  move : ha hPA. clear. hauto q:on use:adequacy.
+Qed.
+
+Lemma ST_App n Γ (b a : PTm n) A B :
+  Γ ⊨ b ∈ PBind PPi A B ->
+  Γ ⊨ a ∈ A ->
+  Γ ⊨ PApp b a ∈ subst_PTm (scons a VarPTm) B.
+Proof.
+  move => hf hb ρ hρ.
+  move /(_ ρ hρ) : hf; intros (i & PPi & hPi & hf).
+  move /(_ ρ hρ) : hb; intros (j & PA & hPA & hb).
+  simpl in hPi.
+  move /InterpUniv_Bind_inv_nopf : hPi. intros (PA0 & hPA0 & hTot & ?). subst.
+  have ? : PA0 = PA by eauto using InterpUniv_Functional'. subst.
+  move  : hf (hb). move/[apply].
+  move : hTot hb. move/[apply].
+  asimpl. hauto lq:on.
+Qed.
+
+Lemma ST_Pair n Γ (a b : PTm n) 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.
+Proof.
+  move /SemWt_Univ => + ha hb ρ hρ.
+  move /(_ _ hρ) => [PPi hPPi].
+  exists i, PPi. split => //.
+  simpl in hPPi.
+  move /InterpUniv_Bind_inv_nopf : hPPi.
+  move => [PA [hPA [hTot ?]]]. subst=>/=.
+  rewrite /SumSpace. right.
+  exists (subst_PTm ρ a), (subst_PTm ρ b).
+  split.
+  - apply rtc_refl.
+  - move /ha : (hρ){ha}.
+    move => [m][PA0][h0]h1.
+    move /hb : (hρ){hb}.
+    move => [k][PB][h2]h3.
+    have ? : PA0 = PA by eauto using InterpUniv_Functional'. subst.
+    split => // PB0.
+    move : h2. asimpl => *.
+    have ? : PB0 = PB by eauto using InterpUniv_Functional'. by subst.
+Qed.
+
+Lemma N_Projs n p (a b : PTm n) :
+  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 :
+  Γ ⊨ a ∈ PBind PSig A B ->
+  Γ ⊨ PProj PL a ∈ A.
+Proof.
+  move => h ρ /[dup]hρ {}/h [m][PA][/= /InterpUniv_Bind_inv_nopf h0]h1.
+  move : h0 => [S][h2][h3]?. subst.
+  move : h1 => /=.
+  rewrite /SumSpace.
+  case.
+  - move => [v [h0 h1]].
+    have {}h0 : rtc TRedSN (PProj PL (subst_PTm ρ a)) (PProj PL v) by hauto lq:on use:N_Projs.
+    have {}h1 : SNe (PProj PL v) by hauto lq:on ctrs:SNe.
+    hauto q:on use:InterpUniv_back_closs,adequacy.
+  - move => [a0 [b0 [h4 [h5 h6]]]].
+    exists m, S. split => //=.
+    have {}h4 : rtc TRedSN (PProj PL (subst_PTm ρ a)) (PProj PL (PPair a0 b0)) by eauto using N_Projs.
+    have ? : rtc TRedSN (PProj PL (PPair a0 b0)) a0 by hauto q:on ctrs:rtc, TRedSN use:adequacy.
+    have : rtc TRedSN (PProj PL (subst_PTm ρ a)) a0 by hauto q:on ctrs:rtc use:@relations.rtc_r.
+    move => h.
+    apply : InterpUniv_back_closs; eauto.
+Qed.
+
+Lemma ST_Proj2 n Γ (a : PTm n) A B :
+  Γ ⊨ a ∈ PBind PSig A B ->
+  Γ ⊨ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B.
+Proof.
+  move => h ρ hρ.
+  move : (hρ) => {}/h [m][PA][/= /InterpUniv_Bind_inv_nopf h0]h1.
+  move : h0 => [S][h2][h3]?. subst.
+  move : h1 => /=.
+  rewrite /SumSpace.
+  case.
+  - move => h.
+    move : h => [v [h0 h1]].
+    have hp : forall p, SNe (PProj p v) by hauto lq:on ctrs:SNe.
+    have hp' : forall p, rtc TRedSN (PProj p(subst_PTm ρ a)) (PProj p v) by eauto using N_Projs.
+    have hp0 := hp PL. have hp1 := hp PR => {hp}.
+    have hp0' := hp' PL. have hp1' := hp' PR => {hp'}.
+    have : S (PProj PL (subst_PTm ρ a)). apply : InterpUniv_back_closs; eauto. hauto q:on use:adequacy.
+    move /h3 => [PB]. asimpl => hPB.
+    do 2 eexists. split; eauto.
+    apply : InterpUniv_back_closs; eauto. hauto q:on use:adequacy.
+  - move => [a0 [b0 [h4 [h5 h6]]]].
+    have h3_dup := h3.
+    specialize h3 with (1 := h5).
+    move : h3 => [PB hPB].
+    have hr : forall p, rtc TRedSN (PProj p (subst_PTm ρ a)) (PProj p (PPair a0 b0)) by hauto l:on use: N_Projs.
+    have hSN : SN a0 by move : h5 h2; clear; hauto q:on use:adequacy.
+    have hSN' : SN b0 by hauto q:on use:adequacy.
+    have hrl : TRedSN (PProj PL (PPair a0 b0)) a0 by hauto lq:on ctrs:TRedSN.
+    have hrr : TRedSN (PProj PR (PPair a0 b0)) b0 by hauto lq:on ctrs:TRedSN.
+    exists m, PB.
+    asimpl. split.
+    + have hr' : rtc TRedSN (PProj PL (subst_PTm ρ a)) a0 by hauto l:on use:@relations.rtc_r.
+      have : S (PProj PL (subst_PTm ρ a)) by hauto lq:on use:InterpUniv_back_closs.
+      move => {}/h3_dup.
+      move => [PB0]. asimpl => hPB0.
+      suff : PB = PB0 by congruence.
+      move : hPB. asimpl => hPB.
+      suff : DJoin.R (subst_PTm (scons (PProj PL (subst_PTm ρ a)) ρ) B) (subst_PTm (scons a0 ρ) B).
+      move : InterpUniv_Join hPB0 hPB; repeat move/[apply]. done.
+      suff : BJoin.R (subst_PTm (scons (PProj PL (subst_PTm ρ a)) ρ) B) (subst_PTm (scons a0 ρ) B)
+        by hauto q:on use:DJoin.FromBJoin.
+      have : BJoin.R (PApp (PAbs (subst_PTm (up_PTm_PTm ρ) B)) (PProj PL (subst_PTm ρ a)))
+               (subst_PTm (scons (PProj PL (subst_PTm ρ a)) ρ) B).
+      eexists. split. apply relations.rtc_once. apply RRed.AppAbs.
+      asimpl. apply rtc_refl.
+      have /BJoin.symmetric : BJoin.R (PApp (PAbs (subst_PTm (up_PTm_PTm ρ)B)) a0)
+                                (subst_PTm (scons a0 ρ) B).
+      eexists. split. apply relations.rtc_once. apply RRed.AppAbs.
+      asimpl. apply rtc_refl.
+      suff : BJoin.R (PApp (PAbs (subst_PTm (up_PTm_PTm ρ) B)) (PProj PL (subst_PTm ρ a)))
+               (PApp (PAbs (subst_PTm (up_PTm_PTm ρ)B)) a0) by eauto using BJoin.transitive, BJoin.symmetric.
+      apply BJoin.AppCong. apply BJoin.refl.
+      move /RReds.FromRedSNs : hr'.
+      hauto lq:on ctrs:rtc unfold:BJoin.R.
+    + hauto lq:on use:@relations.rtc_r, InterpUniv_back_closs.
+Qed.

From c24cc7c9b0c65dd705ec1332837b0ebf96f85a44 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 00:08:02 -0500
Subject: [PATCH 035/210] Finish ST Conv

---
 theories/fp_red.v | 46 ++++++++++++++++++++++++++++++++++++++++++++++
 theories/logrel.v | 15 +++++++++++++++
 2 files changed, 61 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5af61f4..e22e759 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -107,6 +107,10 @@ Module EPar.
     all : hauto lq:on ctrs:R use:morphing_up.
   Qed.
 
+  Lemma substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Proof. hauto l:on use:morphing, refl. Qed.
+
 End EPar.
 
 Inductive SNe {n} : PTm n -> Prop :=
@@ -572,6 +576,15 @@ Module RRed.
     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) :
+    R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
+  Proof.
+    move => h. move : m ρ. elim : n a b / h => n.
+    move => a b m ρ /=.
+    eapply AppAbs'; eauto; cycle 1. by asimpl.
+    all : hauto lq:on ctrs:R.
+  Qed.
+
 End RRed.
 
 Module RPar.
@@ -1482,6 +1495,15 @@ Module ERed.
       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).
+  Proof.
+    move => h. move : m ρ. elim : n a b /h => n.
+    move => a m ρ /=.
+    eapply AppEta'; eauto. by asimpl.
+    all : hauto lq:on ctrs:R.
+  Qed.
+
 End ERed.
 
 Module EReds.
@@ -1546,6 +1568,12 @@ Module EReds.
     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).
+  Proof.
+    induction 1; hauto lq:on ctrs:rtc use:ERed.substing.
+  Qed.
+
 End EReds.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -1632,6 +1660,12 @@ Module RERed.
     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).
+  Proof.
+    hauto q:on use:ToBetaEta, FromBeta, FromEta, RRed.substing, ERed.substing.
+  Qed.
+
 End RERed.
 
 Module REReds.
@@ -1716,6 +1750,12 @@ Module REReds.
     hauto lq:on rew:off ctrs:rtc 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).
+  Proof.
+    induction 1; hauto lq:on ctrs:rtc use:RERed.substing.
+  Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2207,4 +2247,10 @@ Module DJoin.
     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).
+  Proof.
+    hauto lq:on rew:off ctrs:rtc unfold:R use:REReds.substing.
+  Qed.
+
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index e88449b..5377dfb 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -735,3 +735,18 @@ Proof.
       hauto lq:on ctrs:rtc unfold:BJoin.R.
     + hauto lq:on use:@relations.rtc_r, InterpUniv_back_closs.
 Qed.
+
+Lemma ST_Conv n Γ (a : PTm n) A B i :
+  Γ ⊨ a ∈ A ->
+  Γ ⊨ B ∈ PUniv i ->
+  DJoin.R A B ->
+  Γ ⊨ a ∈ B.
+Proof.
+  move => ha /SemWt_Univ h h0.
+  move => ρ hρ.
+  have {}h0 : DJoin.R (subst_PTm ρ A) (subst_PTm ρ B) by eauto using DJoin.substing.
+  move /ha : (hρ){ha} => [m [PA [h1 h2]]].
+  move /h : (hρ){h} => [S hS].
+  have ? : PA = S by eauto using InterpUniv_Join'. subst.
+  eauto.
+Qed.

From fecac84977f7550cf394c3f1b4206479ff25ff65 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 13:21:30 -0500
Subject: [PATCH 036/210] Set up interpretation for typed equality

---
 theories/fp_red.v |  6 +++++
 theories/logrel.v | 57 +++++++++++++++++++++++++++++++++++++++++++++++
 2 files changed, 63 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index e22e759..a650e1d 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2176,6 +2176,12 @@ Module DJoin.
     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 :
+    R A0 A1 ->
+    R B0 B1 ->
+    R (PBind p A0 B0) (PBind p A1 B1).
+  Proof. hauto q:on use:REReds.BindCong. Qed.
+
   Lemma FromRedSNs n (a b : PTm n) :
     rtc TRedSN a b ->
     R a b.
diff --git a/theories/logrel.v b/theories/logrel.v
index 5377dfb..2ed3a21 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -438,9 +438,29 @@ Qed.
 Definition ρ_ok {n} (Γ : fin n -> PTm n) (ρ : fin n -> PTm 0) := forall i k PA,
     ⟦ subst_PTm ρ (Γ i) ⟧ k ↘ PA -> PA (ρ i).
 
+Definition ρ_eq {n} (Γ : fin n -> PTm n) (ρ0 ρ1 : fin n -> PTm 0) := forall i,
+    DJoin.R (subst_PTm ρ0 (Γ i)) (subst_PTm ρ1 (Γ i)).
+
 Definition SemWt {n} Γ (a A : PTm n) := 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).
+Notation "Γ ⊨ a ≡ b ∈ A" := (SemEq Γ a b A) (at level 70).
+
+Lemma SemEq_SemWt n Γ (a b A : PTm n) : Γ ⊨ a ≡ b ∈ A -> Γ ⊨ a ∈ A /\ Γ ⊨ b ∈ A /\ DJoin.R a b.
+Proof. hauto lq:on rew:off unfold:SemEq, SemWt. Qed.
+
+Lemma SemWt_SemEq n Γ (a b A : PTm n) : Γ ⊨ a ∈ A -> Γ ⊨ b ∈ A -> (DJoin.R a b) -> Γ ⊨ a ≡ b ∈ A.
+Proof.
+  move => ha hb heq. split => //= ρ hρ.
+  have {}/ha := hρ.
+  have {}/hb := hρ.
+  move => [k][PA][hPA]hpb.
+  move => [k0][PA0][hPA0]hpa.
+  have : PA = PA0 by hauto l:on use:InterpUniv_Functional'.
+  hauto lq:on.
+Qed.
+
 (* Semantic context wellformedness *)
 Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ ⊨ Γ i ∈ PUniv j.
 Notation "⊨ Γ" := (SemWff Γ) (at level 70).
@@ -750,3 +770,40 @@ Proof.
   have ? : PA = S by eauto using InterpUniv_Join'. subst.
   eauto.
 Qed.
+
+Lemma SE_Refl n Γ (a : PTm n) 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 :
+  Γ ⊨ a ≡ b ∈ A ->
+  Γ ⊨ b ≡ a ∈ A.
+Proof. hauto q:on unfold:SemEq. Qed.
+
+Lemma SE_Transitive n Γ (a b c : PTm n) A :
+  Γ ⊨ a ≡ b ∈ A ->
+  Γ ⊨ b ≡ c ∈ A ->
+  Γ ⊨ a ≡ c ∈ A.
+Proof.
+  move => ha hb.
+  apply SemEq_SemWt in ha, hb.
+  have ? : SN b by hauto l:on use:SemWt_SN.
+  apply SemWt_SemEq; try tauto.
+  hauto l:on use:DJoin.transitive.
+Qed.
+
+Lemma SE_Bind n Γ i j p (A0 A1 : PTm n) B0 B1 :
+  Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv j ->
+  Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv (max i j).
+Proof.
+  move => hA hB.
+  apply SemEq_SemWt in hA, hB.
+  apply SemWt_SemEq.
+  hauto l:on use:ST_Bind.
+  apply ST_Bind; last by hauto l:on use:DJoin.BindCong. tauto.
+  move => ρ hρ.
+  suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
+
+  hauto l:on use:DJoin.BindCong.

From 286ceeed39cc0ee5884c1a5d0651a9d16a1eabe1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 13:53:01 -0500
Subject: [PATCH 037/210] Need to extend the notion of semwt to cover
 arbitrarily scoped terms

---
 theories/logrel.v | 20 ++++++++++++++++----
 1 file changed, 16 insertions(+), 4 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 2ed3a21..5fa16a1 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -800,10 +800,22 @@ Lemma SE_Bind n Γ i j p (A0 A1 : PTm n) B0 B1 :
 Proof.
   move => hA hB.
   apply SemEq_SemWt in hA, hB.
-  apply SemWt_SemEq.
+  apply SemWt_SemEq; last by hauto l:on use:DJoin.BindCong.
   hauto l:on use:ST_Bind.
-  apply ST_Bind; last by hauto l:on use:DJoin.BindCong. tauto.
+  (* rewrite SemWt_Univ. *)
+  (* move => ρ hρ /=. *)
+
+
+  apply ST_Bind; first by tauto.
   move => ρ hρ.
   suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
-
-  hauto l:on use:DJoin.BindCong.
+  rewrite /ρ_ok in hρ *.
+  move => j0 k PA.
+  destruct j0 as [j0|].
+  asimpl.
+  move /(_ (Some j0) k PA) : hρ. by asimpl.
+  asimpl.
+  move /(_ None k PA) : (hρ). asimpl.
+  have /SemWt_Univ h : Γ ⊨ A1 ∈ PUniv i by tauto.
+  have /h := hρ.
+  suff : DJoin.R (subst_PTm (funcomp ρ shift) A1) (subst_PTm (funcomp ρ shift) A0) by best use:InterpUniv_Join.

From db911cff364cfd45037d32d4b39f52640818cee0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 14:31:42 -0500
Subject: [PATCH 038/210] Finish the context equality lemma

---
 theories/logrel.v | 50 ++++++++++++++++++++++++++++++++++++-----------
 1 file changed, 39 insertions(+), 11 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 5fa16a1..ee03604 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -438,8 +438,7 @@ Qed.
 Definition ρ_ok {n} (Γ : fin n -> PTm n) (ρ : fin n -> PTm 0) := forall i k PA,
     ⟦ subst_PTm ρ (Γ i) ⟧ k ↘ PA -> PA (ρ i).
 
-Definition ρ_eq {n} (Γ : fin n -> PTm n) (ρ0 ρ1 : fin n -> PTm 0) := forall i,
-    DJoin.R (subst_PTm ρ0 (Γ i)) (subst_PTm ρ1 (Γ i)).
+Definition Γ_eq {n} (Γ Δ : fin n -> PTm n)  := forall i, DJoin.R (Γ i) (Δ i).
 
 Definition SemWt {n} Γ (a A : PTm n) := forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a).
 Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70).
@@ -489,6 +488,13 @@ Proof.
     by subst.
 Qed.
 
+Lemma ρ_ok_cons' n i (Γ : fin n -> PTm n) ρ a PA A  Δ :
+  Δ = (funcomp (ren_PTm shift) (scons A Γ)) ->
+  ⟦ subst_PTm ρ A ⟧ i ↘ PA -> PA a ->
+  ρ_ok Γ ρ ->
+  ρ_ok Δ (scons a ρ).
+Proof. move => ->. apply ρ_ok_cons. Qed.
+
 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).
 
@@ -793,29 +799,51 @@ Proof.
   hauto l:on use:DJoin.transitive.
 Qed.
 
+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 SE_Bind n Γ i j p (A0 A1 : PTm n) B0 B1 :
+  ⊨ Γ ->
   Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
   funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv j ->
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv (max i j).
 Proof.
-  move => hA hB.
+  move => hΓ 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.
-  (* rewrite SemWt_Univ. *)
-  (* move => ρ hρ /=. *)
-
 
   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.
-  rewrite /ρ_ok in hρ *.
+  best use:
   move => j0 k PA.
   destruct j0 as [j0|].
   asimpl.
   move /(_ (Some j0) k PA) : hρ. by asimpl.
   asimpl.
-  move /(_ None k PA) : (hρ). asimpl.
-  have /SemWt_Univ h : Γ ⊨ A1 ∈ PUniv i by tauto.
-  have /h := hρ.
-  suff : DJoin.R (subst_PTm (funcomp ρ shift) A1) (subst_PTm (funcomp ρ shift) A0) by best use:InterpUniv_Join.
+  have h : Γ ⊨ A1 ∈ PUniv i by tauto.
+  rewrite /SemWff in hΓ'.
+  move /(_ None) : hΓ' => [l h'].
+  rewrite SemWt_Univ in h'.
+  have {}/h' := hρ.
+  move => [PA'].
+  asimpl. move => h0 h1.
+  move /(_ None l PA) : (hρ). asimpl.
+  suff : PA = PA' by qauto l:on.
+  move : InterpUniv_Join' h1 h0; repeat move/[apply].
+  apply. apply DJoin.substing. tauto.
+Qed.

From 1258ac263c9e467d6865395d9e5f404fb29a2db2 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 14:37:25 -0500
Subject: [PATCH 039/210] Add structural rule for ctx equivalence

---
 theories/logrel.v | 26 +++++++++++++++++++++++++-
 1 file changed, 25 insertions(+), 1 deletion(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index ee03604..f57c323 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -814,6 +814,28 @@ Proof.
   hauto l:on use: DJoin.substing.
 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 :
   ⊨ Γ ->
   Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
@@ -824,11 +846,13 @@ Proof.
   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.
+
+
   best use:
   move => j0 k PA.
   destruct j0 as [j0|].

From 435c0e037e5cede64c768e8fdc9625e88cfb7c80 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 14:37:56 -0500
Subject: [PATCH 040/210] Finish (simplified) SE_Bind

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

diff --git a/theories/logrel.v b/theories/logrel.v
index f57c323..9c39125 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -851,23 +851,5 @@ Proof.
   move => ρ hρ.
   suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
   move : Γ_eq_ρ_ok hΓ' hρ; repeat move/[apply]. apply.
-
-
-  best use:
-  move => j0 k PA.
-  destruct j0 as [j0|].
-  asimpl.
-  move /(_ (Some j0) k PA) : hρ. by asimpl.
-  asimpl.
-  have h : Γ ⊨ A1 ∈ PUniv i by tauto.
-  rewrite /SemWff in hΓ'.
-  move /(_ None) : hΓ' => [l h'].
-  rewrite SemWt_Univ in h'.
-  have {}/h' := hρ.
-  move => [PA'].
-  asimpl. move => h0 h1.
-  move /(_ None l PA) : (hρ). asimpl.
-  suff : PA = PA' by qauto l:on.
-  move : InterpUniv_Join' h1 h0; repeat move/[apply].
-  apply. apply DJoin.substing. tauto.
+  hauto lq:on use:Γ_eq_cons'.
 Qed.

From 733e86c6118ac8c091e6fb33ff266a5a4e446ad1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 15:20:40 -0500
Subject: [PATCH 041/210] Finish SE_Pair

---
 theories/fp_red.v | 26 +++++++++++++++++
 theories/logrel.v | 74 +++++++++++++++++++++++++++++++++--------------
 2 files changed, 78 insertions(+), 22 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index a650e1d..9afad56 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1756,6 +1756,24 @@ Module REReds.
     induction 1; hauto lq:on ctrs:rtc use:RERed.substing.
   Qed.
 
+
+  Lemma cong_up n m (ρ0 ρ1 : fin n -> PTm m) :
+    (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
+    (forall i, rtc RERed.R (up_PTm_PTm ρ0 i) (up_PTm_PTm ρ1 i)).
+  Proof. move => h i. destruct i as [i|].
+         simpl. rewrite /funcomp.
+         substify. by apply substing.
+         apply rtc_refl.
+  Qed.
+
+  Lemma cong n m (a : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
+    (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
+    rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 a).
+  Proof.
+    move : m ρ0 ρ1. elim : n / a;
+      eauto using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl.
+  Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2259,4 +2277,12 @@ Module DJoin.
     hauto lq:on rew:off ctrs:rtc unfold:R use:REReds.substing.
   Qed.
 
+  Lemma cong n m (a : PTm (S n)) c d (ρ : fin n -> PTm m) :
+    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.
+  Qed.
+
 End DJoin.
diff --git a/theories/logrel.v b/theories/logrel.v
index 9c39125..2c698b2 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -350,14 +350,7 @@ Proof.
     have ? : PA0 = PA by qauto l:on. subst.
     have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'.
     move => a PB PB' ha hPB hPB'. apply : ihPF; eauto.
-    have hj0 : DJoin.R (PAbs B) (PAbs B0) by eauto using DJoin.AbsCong.
-    have {}hj0 : DJoin.R (PApp (PAbs B) a) (PApp (PAbs B0) a) by eauto using DJoin.AppCong, DJoin.refl.
-    have [? ?] : SN (PApp (PAbs B) a) /\ SN (PApp (PAbs B0) a) by
-      hauto lq:on rew:off use:N_Exp, N_β, adequacy.
-    have [? ?] : DJoin.R (PApp (PAbs B0) a) (subst_PTm (scons a VarPTm) B0) /\
-           DJoin.R (subst_PTm (scons a VarPTm) B) (PApp (PAbs B) a)
-      by hauto lq:on ctrs:RRed.R use:DJoin.FromRRed0, DJoin.FromRRed1.
-    eauto using DJoin.transitive.
+    by apply DJoin.substing.
     move => h. extensionality b. apply propositional_extensionality.
     hauto l:on use:bindspace_iff.
   - move => i j jlti ih B PB hPB.
@@ -744,20 +737,8 @@ Proof.
       move : hPB. asimpl => hPB.
       suff : DJoin.R (subst_PTm (scons (PProj PL (subst_PTm ρ a)) ρ) B) (subst_PTm (scons a0 ρ) B).
       move : InterpUniv_Join hPB0 hPB; repeat move/[apply]. done.
-      suff : BJoin.R (subst_PTm (scons (PProj PL (subst_PTm ρ a)) ρ) B) (subst_PTm (scons a0 ρ) B)
-        by hauto q:on use:DJoin.FromBJoin.
-      have : BJoin.R (PApp (PAbs (subst_PTm (up_PTm_PTm ρ) B)) (PProj PL (subst_PTm ρ a)))
-               (subst_PTm (scons (PProj PL (subst_PTm ρ a)) ρ) B).
-      eexists. split. apply relations.rtc_once. apply RRed.AppAbs.
-      asimpl. apply rtc_refl.
-      have /BJoin.symmetric : BJoin.R (PApp (PAbs (subst_PTm (up_PTm_PTm ρ)B)) a0)
-                                (subst_PTm (scons a0 ρ) B).
-      eexists. split. apply relations.rtc_once. apply RRed.AppAbs.
-      asimpl. apply rtc_refl.
-      suff : BJoin.R (PApp (PAbs (subst_PTm (up_PTm_PTm ρ) B)) (PProj PL (subst_PTm ρ a)))
-               (PApp (PAbs (subst_PTm (up_PTm_PTm ρ)B)) a0) by eauto using BJoin.transitive, BJoin.symmetric.
-      apply BJoin.AppCong. apply BJoin.refl.
-      move /RReds.FromRedSNs : hr'.
+      apply DJoin.cong.
+      apply DJoin.FromRedSNs.
       hauto lq:on ctrs:rtc unfold:BJoin.R.
     + hauto lq:on use:@relations.rtc_r, InterpUniv_back_closs.
 Qed.
@@ -853,3 +834,52 @@ Proof.
   move : Γ_eq_ρ_ok hΓ' hρ; repeat move/[apply]. apply.
   hauto lq:on use:Γ_eq_cons'.
 Qed.
+
+Lemma SE_Abs n Γ (a b : PTm (S n)) A B i :
+  Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
+  funcomp (ren_PTm shift) (scons 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 SE_Conv n Γ (a b : PTm n) A B i :
+  Γ ⊨ a ≡ b ∈ A ->
+  Γ ⊨ B ∈ PUniv i ->
+  DJoin.R A B ->
+  Γ ⊨ a ≡ b ∈ B.
+Proof.
+  move /SemEq_SemWt => [ha][hb]he hB hAB.
+  apply SemWt_SemEq; eauto using ST_Conv.
+Qed.
+
+Lemma SBind_inst n Γ p i (A : PTm n) B (a : PTm n) :
+  Γ ⊨ a ∈ A ->
+  Γ ⊨ PBind p A B ∈ PUniv i ->
+  Γ ⊨ subst_PTm (scons a VarPTm) B ∈ PUniv i.
+Proof.
+  move => ha /SemWt_Univ hb.
+  apply SemWt_Univ.
+  move => ρ hρ.
+  have {}/hb := hρ.
+  asimpl. move => /= [S hS].
+  move /InterpUniv_Bind_inv_nopf : hS.
+  move => [PA][hPA][hPF]?. subst.
+  have {}/ha := hρ.
+  move => [k][PA0][hPA0]ha.
+  have ? : PA0 = PA by hauto l:on use:InterpUniv_Functional'. subst.
+  have {}/hPF := ha.
+  move => [PB]. asimpl.
+  hauto lq:on.
+Qed.
+
+Lemma SE_Pair n Γ (a0 a1 b0 b1 : PTm n) 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.
+Proof.
+  move => h /SemEq_SemWt [ha0][ha1]hae /SemEq_SemWt [hb0][hb1]hbe.
+  apply SemWt_SemEq; eauto using ST_Pair, DJoin.PairCong, SBind_inst, DJoin.cong, ST_Conv, ST_Pair.
+Qed.

From aca7ebaf1e886a2dc1639dcbce222a7f7e3736a6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 15:42:35 -0500
Subject: [PATCH 042/210] Finish all the semantic theory for the system with
 type-directed eq

---
 theories/logrel.v | 35 +++++++++++++++++++++++++++++++++++
 1 file changed, 35 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index 2c698b2..c0b2fe7 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -883,3 +883,38 @@ Proof.
   move => h /SemEq_SemWt [ha0][ha1]hae /SemEq_SemWt [hb0][hb1]hbe.
   apply SemWt_SemEq; eauto using ST_Pair, DJoin.PairCong, SBind_inst, DJoin.cong, ST_Conv, ST_Pair.
 Qed.
+
+Lemma SE_Proj1 n Γ (a b : PTm n) A B :
+  Γ ⊨ a ≡ b ∈ PBind PSig A B ->
+  Γ ⊨ PProj PL a ≡ PProj PL b ∈ A.
+Proof.
+  move => /SemEq_SemWt [ha][hb]he.
+  apply SemWt_SemEq; eauto using DJoin.ProjCong, ST_Proj1.
+Qed.
+
+Lemma SE_Proj2 n Γ i (a b : PTm n) 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.
+Proof.
+  move => hS.
+  move => /SemEq_SemWt [ha][hb]he.
+  apply SemWt_SemEq; eauto using DJoin.ProjCong, ST_Proj2.
+  have h : Γ ⊨ PProj PR b ∈ subst_PTm (scons (PProj PL b) VarPTm) B by eauto using ST_Proj2.
+  apply : ST_Conv. apply h.
+  apply : SBind_inst. eauto using ST_Proj1.
+  eauto.
+  hauto lq:on use: DJoin.cong,  DJoin.ProjCong.
+Qed.
+
+Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) 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.
+Proof.
+  move => hPi.
+  move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha.
+  apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App.
+  apply : ST_Conv; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
+Qed.

From c25bac311c41818e4e54d0ebfb00457c43ebd54a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 16:02:03 -0500
Subject: [PATCH 043/210] Add the first-component inversion lemma for pi and
 sigma

---
 theories/logrel.v | 46 ++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 46 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index c0b2fe7..2579967 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -918,3 +918,49 @@ Proof.
   apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App.
   apply : ST_Conv; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
 Qed.
+
+Lemma SBind_inv1 n Γ i p (A : PTm n) 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 SBind_inv2 n Γ i p (A : PTm n) B :
+  Γ ⊨ PBind p A B ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv i.
+  move /SemWt_Univ => h. apply SemWt_Univ.
+  move => ρ hρ.
+  set ρ' := funcomp ρ shift.
+  move /(_ ρ') : h.
+  have : ρ_ok Γ ρ'.
+  move => k.
+  subst ρ'.
+  rewrite /ρ_ok in hρ.
+  move => l PA. move => /(_ (shift k) l PA) in hρ.
+  move : hρ. asimpl. rewrite /funcomp. hauto lq:on.
+  move /[swap]/[apply].
+  move => /= [S hS].
+  move /InterpUniv_Bind_inv_nopf : hS.
+  move => [PA][hPA][hPF]?. subst.
+  subst ρ'.
+
+Qed.
+
+
+
+Lemma Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
+  Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
+  Γ ⊨ A0 ≡ A1 ∈ PUniv i.
+Proof.
+  move /SemEq_SemWt => [h0][h1]he.
+  apply SemWt_SemEq; eauto using SBind_inv.
+  move /DJoin.bind_inj : he. tauto.
+Qed.
+
+Lemma Bind_Proj2' n Γ p (A0 A1 : PTm n) B0 B1 i :
+  Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv i.
+Proof.
+  move /SemEq_SemWt => [h0][h1]he.
+  apply SemWt_SemEq; last by hauto l:on use:DJoin.bind_inj.

From e7a26e6bd6724248d59aa5342d13b15a5b2269ef Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 16:20:56 -0500
Subject: [PATCH 044/210] Finish all Bind proj lemmas

---
 theories/logrel.v | 49 ++++++++++++++++++++---------------------------
 1 file changed, 21 insertions(+), 28 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 2579967..429bcd4 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -926,41 +926,34 @@ Lemma SBind_inv1 n Γ i p (A : PTm n) B :
   hauto lq:on rew:off  use:InterpUniv_Bind_inv.
 Qed.
 
-Lemma SBind_inv2 n Γ i p (A : PTm n) B :
-  Γ ⊨ PBind p A B ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv i.
-  move /SemWt_Univ => h. apply SemWt_Univ.
-  move => ρ hρ.
-  set ρ' := funcomp ρ shift.
-  move /(_ ρ') : h.
-  have : ρ_ok Γ ρ'.
-  move => k.
-  subst ρ'.
-  rewrite /ρ_ok in hρ.
-  move => l PA. move => /(_ (shift k) l PA) in hρ.
-  move : hρ. asimpl. rewrite /funcomp. hauto lq:on.
-  move /[swap]/[apply].
-  move => /= [S hS].
-  move /InterpUniv_Bind_inv_nopf : hS.
-  move => [PA][hPA][hPF]?. subst.
-  subst ρ'.
-
-Qed.
-
-
-
 Lemma Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
   Γ ⊨ A0 ≡ A1 ∈ PUniv i.
 Proof.
   move /SemEq_SemWt => [h0][h1]he.
-  apply SemWt_SemEq; eauto using SBind_inv.
+  apply SemWt_SemEq; eauto using SBind_inv1.
   move /DJoin.bind_inj : he. tauto.
 Qed.
 
-Lemma Bind_Proj2' n Γ p (A0 A1 : PTm n) B0 B1 i :
+Lemma Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv i.
+  Γ ⊨ a0 ≡ a1 ∈ A0 ->
+  Γ ⊨ subst_PTm (scons a0 VarPTm) B0 ≡ subst_PTm (scons a1 VarPTm) B1 ∈ PUniv i.
 Proof.
-  move /SemEq_SemWt => [h0][h1]he.
-  apply SemWt_SemEq; last by hauto l:on use:DJoin.bind_inj.
+  move /SemEq_SemWt => [hA][hB]/DJoin.bind_inj [_ [h1 h2]] /SemEq_SemWt [ha0][ha1]ha.
+  move : h2 => [B [h2 h2']].
+  move : ha => [c [ha ha']].
+  apply SemWt_SemEq; eauto using SBind_inst.
+  apply SBind_inst with (p := p) (A := A1); eauto.
+  apply : ST_Conv; eauto. hauto l:on use:SBind_inv1.
+  rewrite /DJoin.R.
+  eexists. split.
+  apply : relations.rtc_transitive.
+  apply REReds.substing. apply h2.
+  apply REReds.cong.
+  have : forall i0, rtc RERed.R (scons a0 VarPTm i0) (scons c VarPTm i0) by hauto q:on ctrs:rtc inv:option.
+  apply.
+  apply : relations.rtc_transitive. apply REReds.substing; eauto.
+  apply REReds.cong.
+  hauto lq:on ctrs:rtc inv:option.
+Qed.

From 82e21196c2d2de9dad51d70de8b6498022bb27ca Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 16:31:50 -0500
Subject: [PATCH 045/210] Minor

---
 theories/logrel.v | 5 +++++
 1 file changed, 5 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index 429bcd4..647974f 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -539,6 +539,11 @@ Proof.
   move {h}. hauto l:on use:sn_unmorphing.
 Qed.
 
+Lemma SemEq_SN_Join n Γ (a b A : PTm n) :
+  Γ ⊨ a ≡ b ∈ A ->
+  SN a /\ SN b /\ SN A /\ DJoin.R a b.
+Proof. hauto l:on use:SemEq_SemWt, SemWt_SN. Qed.
+
 Lemma SemWt_Univ n Γ (A : PTm n) i  :
   Γ ⊨ A ∈ PUniv i <->
   forall ρ, ρ_ok Γ ρ -> exists S, ⟦ subst_PTm ρ A ⟧ i ↘ S.

From 10f339c5b6e769157ec2a8c6e433c7f52efbb450 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 17:00:04 -0500
Subject: [PATCH 046/210] Add some syntactic typing rules

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

diff --git a/theories/typing.v b/theories/typing.v
new file mode 100644
index 0000000..bd0a0f4
--- /dev/null
+++ b/theories/typing.v
@@ -0,0 +1,49 @@
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
+
+Reserved Notation "Γ ⊢ a ∈ A" (at level 70).
+Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70).
+Reserved Notation "⊢ Γ" (at level 70).
+Inductive Wt {n} : (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
+| T_Var Γ (i : fin n) :
+  ⊢ Γ ->
+  Γ ⊢ VarPTm i ∈ Γ i
+
+| T_Bind Γ i j p (A : PTm n) (B : PTm (S n)) :
+  Γ ⊢ A ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv j ->
+  Γ ⊢ PBind p A B ∈ PUniv (max i j)
+
+| T_Abs Γ (a : PTm (S n)) A B i :
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
+  Γ ⊢ PAbs a ∈ PBind PPi A B
+
+| T_App Γ (b a : PTm n) A B :
+  Γ ⊢ b ∈ PBind PPi A B ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ PApp b a ∈ subst_PTm (scons a VarPTm) B
+
+| T_Pair Γ (a b : PTm n) A B i :
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ PPair a b ∈ PBind PSig A B
+
+| T_Proj1 Γ (a : PTm n) A B :
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ PProj PL a ∈ A
+
+| T_Proj2 Γ (a : PTm n) A B :
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B
+
+| T_Conv Γ (a : PTm n) A B i :
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ A ≡ B ∈ PUniv i ->
+  Γ ⊢ a ∈ B
+
+with Eq {n} : (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
+
+with Wff {n} : (fin n -> PTm n) -> Prop :=
+  where
+  "Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A).

From 8d92f19d74921678e3a760dc05436fc8e6d23ab9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 17:47:34 -0500
Subject: [PATCH 047/210] Add all(?) typing rules

---
 theories/typing.v | 93 ++++++++++++++++++++++++++++++++++++++++-------
 1 file changed, 80 insertions(+), 13 deletions(-)

diff --git a/theories/typing.v b/theories/typing.v
index bd0a0f4..b86fcf6 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -3,47 +3,114 @@ Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
 Reserved Notation "Γ ⊢ a ∈ A" (at level 70).
 Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70).
 Reserved Notation "⊢ Γ" (at level 70).
-Inductive Wt {n} : (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
-| T_Var Γ (i : fin n) :
+Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
+| T_Var n Γ (i : fin n) :
   ⊢ Γ ->
   Γ ⊢ VarPTm i ∈ Γ i
 
-| T_Bind Γ i j p (A : PTm n) (B : PTm (S n)) :
+| T_Bind n Γ i j p (A : PTm n) (B : PTm (S n)) :
   Γ ⊢ A ∈ PUniv i ->
   funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv j ->
   Γ ⊢ PBind p A B ∈ PUniv (max i j)
 
-| T_Abs Γ (a : PTm (S n)) A B i :
+| T_Abs n Γ (a : PTm (S n)) A B i :
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
   funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B ->
   Γ ⊢ PAbs a ∈ PBind PPi A B
 
-| T_App Γ (b a : PTm n) A B :
+| T_App n Γ (b a : PTm n) A B :
   Γ ⊢ b ∈ PBind PPi A B ->
   Γ ⊢ a ∈ A ->
   Γ ⊢ PApp b a ∈ subst_PTm (scons a VarPTm) B
 
-| T_Pair Γ (a b : PTm n) A B i :
+| T_Pair n Γ (a b : PTm n) A B i :
   Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊢ a ∈ A ->
   Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B ->
   Γ ⊢ PPair a b ∈ PBind PSig A B
 
-| T_Proj1 Γ (a : PTm n) A B :
+| T_Proj1 n Γ (a : PTm n) A B :
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PL a ∈ A
 
-| T_Proj2 Γ (a : PTm n) A B :
+| T_Proj2 n Γ (a : PTm n) A B :
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B
 
-| T_Conv Γ (a : PTm n) A B i :
+| T_Conv n Γ (a : PTm n) A B i :
   Γ ⊢ a ∈ A ->
   Γ ⊢ A ≡ B ∈ PUniv i ->
   Γ ⊢ a ∈ B
 
-with Eq {n} : (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
+with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
+| E_Refl n Γ (a : PTm n) A :
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ a ≡ a ∈ A
 
-with Wff {n} : (fin n -> PTm n) -> Prop :=
-  where
-  "Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A).
+| E_Symmetric n Γ (a b : PTm n) A :
+  Γ ⊢ a ≡ b ∈ A ->
+  Γ ⊢ b ≡ a ∈ A
+
+| E_Transitive n Γ (a b c : PTm n) A :
+  Γ ⊢ a ≡ b ∈ A ->
+  Γ ⊢ b ≡ c ∈ A ->
+  Γ ⊢ a ≡ c ∈ A
+
+| E_Bind n Γ i j p (A0 A1 : PTm n) B0 B1 :
+  ⊢ Γ ->
+  Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv j ->
+  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv (max i j)
+
+| E_Abs n Γ (a b : PTm (S n)) A B i :
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ≡ b ∈ B ->
+  Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B
+
+| E_App n Γ i (b0 b1 a0 a1 : PTm n) 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 :
+  Γ ⊢ 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 :
+  Γ ⊢ a ≡ b ∈ PBind PSig A B ->
+  Γ ⊢ PProj PL a ≡ PProj PL b ∈ A
+
+| E_Proj2 n Γ i (a b : PTm n) 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_Conv n Γ (a b : PTm n) A B i :
+  Γ ⊢ a ≡ b ∈ A ->
+  Γ ⊢ B ∈ PUniv i ->
+  Γ ⊢ A ≡ B ∈ PUniv i ->
+  Γ ⊢ a ≡ b ∈ B
+
+| E_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
+  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
+  Γ ⊢ A0 ≡ A1 ∈ PUniv i
+
+| E_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
+  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
+  Γ ⊢ a0 ≡ a1 ∈ A0 ->
+  Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≡ subst_PTm (scons a1 VarPTm) B1 ∈ PUniv i
+
+with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
+| Wff_Nil :
+  ⊢ null
+| Wff_Cons n Γ (A : PTm n) i :
+  ⊢ Γ ->
+  Γ ⊢ A ∈ PUniv i ->
+  (* -------------- *)
+  ⊢ funcomp (ren_PTm shift) (scons A Γ)
+
+where
+"Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A).

From 399c97fa8218a19eebbfb7b591bf81923f644387 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 17:58:38 -0500
Subject: [PATCH 048/210] Add the semantic well-typedness rules to the hint db

---
 theories/logrel.v | 37 ++++++++++++++++++++++++++++---------
 theories/typing.v |  2 +-
 2 files changed, 29 insertions(+), 10 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 647974f..e9127c0 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -748,7 +748,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' n Γ (a : PTm n) A B i :
   Γ ⊨ a ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
   DJoin.R A B ->
@@ -763,6 +763,12 @@ Proof.
   eauto.
 Qed.
 
+Lemma ST_Conv n Γ (a : PTm n) A B i :
+  Γ ⊨ a ∈ A ->
+  Γ ⊨ A ≡ B ∈ PUniv i ->
+  Γ ⊨ a ∈ B.
+Proof. hauto l:on use:ST_Conv', SemEq_SemWt. Qed.
+
 Lemma SE_Refl n Γ (a : PTm n) A :
   Γ ⊨ a ∈ A ->
   Γ ⊨ a ≡ a ∈ A.
@@ -849,14 +855,23 @@ Proof.
   apply SemWt_SemEq; eauto using DJoin.AbsCong, ST_Abs.
 Qed.
 
-Lemma SE_Conv n Γ (a b : PTm n) A B i :
+Lemma SE_Conv' n Γ (a b : PTm n) A B i :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
   DJoin.R A B ->
   Γ ⊨ a ≡ b ∈ B.
 Proof.
   move /SemEq_SemWt => [ha][hb]he hB hAB.
-  apply SemWt_SemEq; eauto using ST_Conv.
+  apply SemWt_SemEq; eauto using ST_Conv'.
+Qed.
+
+Lemma SE_Conv n Γ (a b : PTm n) A B i :
+  Γ ⊨ a ≡ b ∈ A ->
+  Γ ⊨ A ≡ B ∈ PUniv i ->
+  Γ ⊨ a ≡ b ∈ B.
+Proof.
+  move => h /SemEq_SemWt [h0][h1]he.
+  eauto using SE_Conv'.
 Qed.
 
 Lemma SBind_inst n Γ p i (A : PTm n) B (a : PTm n) :
@@ -886,7 +901,7 @@ Lemma SE_Pair n Γ (a0 a1 b0 b1 : PTm n) A B i :
   Γ ⊨ PPair a0 b0 ≡ PPair a1 b1 ∈ PBind PSig A B.
 Proof.
   move => h /SemEq_SemWt [ha0][ha1]hae /SemEq_SemWt [hb0][hb1]hbe.
-  apply SemWt_SemEq; eauto using ST_Pair, DJoin.PairCong, SBind_inst, DJoin.cong, ST_Conv, ST_Pair.
+  apply SemWt_SemEq; eauto using ST_Pair, DJoin.PairCong, SBind_inst, DJoin.cong, ST_Conv', ST_Pair.
 Qed.
 
 Lemma SE_Proj1 n Γ (a b : PTm n) A B :
@@ -906,7 +921,7 @@ Proof.
   move => /SemEq_SemWt [ha][hb]he.
   apply SemWt_SemEq; eauto using DJoin.ProjCong, ST_Proj2.
   have h : Γ ⊨ PProj PR b ∈ subst_PTm (scons (PProj PL b) VarPTm) B by eauto using ST_Proj2.
-  apply : ST_Conv. apply h.
+  apply : ST_Conv'. apply h.
   apply : SBind_inst. eauto using ST_Proj1.
   eauto.
   hauto lq:on use: DJoin.cong,  DJoin.ProjCong.
@@ -921,7 +936,7 @@ Proof.
   move => hPi.
   move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha.
   apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App.
-  apply : ST_Conv; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
+  apply : ST_Conv'; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
 Qed.
 
 Lemma SBind_inv1 n Γ i p (A : PTm n) B :
@@ -931,7 +946,7 @@ Lemma SBind_inv1 n Γ i p (A : PTm n) B :
   hauto lq:on rew:off  use:InterpUniv_Bind_inv.
 Qed.
 
-Lemma Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
+Lemma SE_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
   Γ ⊨ A0 ≡ A1 ∈ PUniv i.
 Proof.
@@ -940,7 +955,7 @@ Proof.
   move /DJoin.bind_inj : he. tauto.
 Qed.
 
-Lemma Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
+Lemma SE_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
   Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
   Γ ⊨ a0 ≡ a1 ∈ A0 ->
   Γ ⊨ subst_PTm (scons a0 VarPTm) B0 ≡ subst_PTm (scons a1 VarPTm) B1 ∈ PUniv i.
@@ -950,7 +965,7 @@ Proof.
   move : ha => [c [ha ha']].
   apply SemWt_SemEq; eauto using SBind_inst.
   apply SBind_inst with (p := p) (A := A1); eauto.
-  apply : ST_Conv; eauto. hauto l:on use:SBind_inv1.
+  apply : ST_Conv'; eauto. hauto l:on use:SBind_inv1.
   rewrite /DJoin.R.
   eexists. split.
   apply : relations.rtc_transitive.
@@ -962,3 +977,7 @@ Proof.
   apply REReds.cong.
   hauto lq:on ctrs:rtc inv:option.
 Qed.
+
+#[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Conv
+  SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
+  SE_Conv SE_Bind_Proj1 SE_Bind_Proj2 SemWff_nil SemWff_cons : sem.
diff --git a/theories/typing.v b/theories/typing.v
index b86fcf6..6b8db7a 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -109,7 +109,7 @@ with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
 | Wff_Cons n Γ (A : PTm n) i :
   ⊢ Γ ->
   Γ ⊢ A ∈ PUniv i ->
-  (* -------------- *)
+  (* -------------------------------- *)
   ⊢ funcomp (ren_PTm shift) (scons A Γ)
 
 where

From 6daa275ac87534efa16e7496281cd70985d1de58 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 18:09:19 -0500
Subject: [PATCH 049/210] Prove the fundamental theorem

---
 theories/soundness.v | 11 +++++++++++
 theories/typing.v    |  6 ++++++
 2 files changed, 17 insertions(+)
 create mode 100644 theories/soundness.v

diff --git a/theories/soundness.v b/theories/soundness.v
new file mode 100644
index 0000000..c2a40f0
--- /dev/null
+++ b/theories/soundness.v
@@ -0,0 +1,11 @@
+Require Import Autosubst2.fintype Autosubst2.syntax.
+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).
+  apply wt_mutual; eauto with sem;[idtac].
+  hauto l:on use:SE_Pair.
+Qed.
diff --git a/theories/typing.v b/theories/typing.v
index 6b8db7a..043be19 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -114,3 +114,9 @@ with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
 
 where
 "Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A).
+
+Scheme wf_ind := Induction for Wff Sort Prop
+  with wt_ind := Induction for Wt Sort Prop
+  with eq_ind := Induction for Eq Sort Prop.
+
+Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind.

From 2e2e41cbe6e85de56ce696e4d517bdb2163f0dbb Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 18:15:25 -0500
Subject: [PATCH 050/210] Add missing Univ rule

---
 theories/logrel.v | 10 +++++++++-
 theories/typing.v |  4 ++++
 2 files changed, 13 insertions(+), 1 deletion(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index e9127c0..a172ec1 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -978,6 +978,14 @@ Proof.
   hauto lq:on ctrs:rtc inv:option.
 Qed.
 
-#[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Conv
+Lemma ST_Univ n Γ i j :
+  i < j ->
+  Γ ⊨ PUniv i : PTm n ∈ PUniv j.
+Proof.
+  move => ?.
+  apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ.
+Qed.
+
+#[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
   SE_Conv SE_Bind_Proj1 SE_Bind_Proj2 SemWff_nil SemWff_cons : sem.
diff --git a/theories/typing.v b/theories/typing.v
index 043be19..e2541b5 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -37,6 +37,10 @@ Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B
 
+| T_Univ n Γ i j :
+  i < j ->
+  Γ ⊢ PUniv i : PTm n ∈ PUniv j
+
 | T_Conv n Γ (a : PTm n) A B i :
   Γ ⊢ a ∈ A ->
   Γ ⊢ A ≡ B ∈ PUniv i ->

From 0e5b82b162104717571074583491b1c54db2f133 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 21:40:26 -0500
Subject: [PATCH 051/210] Move projection axioms to the subtyping relation

---
 theories/common.v     |  3 +++
 theories/logrel.v     |  5 +---
 theories/structural.v | 49 +++++++++++++++++++++++++++++++++++
 theories/typing.v     | 60 +++++++++++++++++++++++++++++++------------
 4 files changed, 96 insertions(+), 21 deletions(-)
 create mode 100644 theories/common.v
 create mode 100644 theories/structural.v

diff --git a/theories/common.v b/theories/common.v
new file mode 100644
index 0000000..d345d49
--- /dev/null
+++ b/theories/common.v
@@ -0,0 +1,3 @@
+Require Import Autosubst2.fintype Autosubst2.syntax.
+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).
diff --git a/theories/logrel.v b/theories/logrel.v
index a172ec1..4374527 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1,5 +1,5 @@
 Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
-Require Import fp_red.
+Require Import common fp_red.
 From Hammer Require Import Tactics.
 From Equations Require Import Equations.
 Require Import ssreflect ssrbool.
@@ -488,9 +488,6 @@ Lemma ρ_ok_cons' n i (Γ : fin n -> PTm n) ρ a PA A  Δ :
   ρ_ok Δ (scons a ρ).
 Proof. move => ->. apply ρ_ok_cons. Qed.
 
-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).
-
 Lemma ρ_ok_renaming n m (Γ : fin n -> PTm n) ρ :
   forall (Δ : fin m -> PTm m) ξ,
     renaming_ok Γ Δ ξ ->
diff --git a/theories/structural.v b/theories/structural.v
new file mode 100644
index 0000000..93fa17c
--- /dev/null
+++ b/theories/structural.v
@@ -0,0 +1,49 @@
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing.
+From Hammer Require Import Tactics.
+Require Import ssreflect.
+
+
+Lemma lem :
+  (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 -> ...).
+Proof. apply wt_mutual. ...
+
+
+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 -> ⊢ Γ).
+Proof. apply wt_mutual; eauto. Qed.
+
+#[export]Hint Constructors Wt Wff Eq : wt.
+
+Lemma renaming_up n m (ξ : fin n -> fin m) Δ Γ A :
+  renaming_ok Δ Γ ξ ->
+  renaming_ok (funcomp (ren_PTm shift) (scons (ren_PTm ξ A) Δ)) (funcomp (ren_PTm shift) (scons 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.
+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 Δ Γ ξ ->
+     Δ ⊢ ren_PTm ξ a ∈ ren_PTm ξ A) /\
+  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> forall m Δ (ξ : fin n -> fin m), ⊢ Δ -> renaming_ok Δ Γ ξ ->
+     Δ ⊢ ren_PTm ξ a ≡ ren_PTm ξ b ∈ ren_PTm ξ A).
+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ξ.
+    apply : T_Abs; eauto.
+    apply iha; last by apply renaming_up.
+    econstructor; eauto.
+    by apply renaming_up.
diff --git a/theories/typing.v b/theories/typing.v
index e2541b5..2d7078e 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -2,6 +2,7 @@ Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
 
 Reserved Notation "Γ ⊢ a ∈ A" (at level 70).
 Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70).
+Reserved Notation "Γ ⊢ A ≲ B" (at level 70).
 Reserved Notation "⊢ Γ" (at level 70).
 Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
 | T_Var n Γ (i : fin n) :
@@ -37,13 +38,13 @@ Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   Γ ⊢ a ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ∈ subst_PTm (scons (PProj PL a) VarPTm) B
 
-| T_Univ n Γ i j :
-  i < j ->
-  Γ ⊢ PUniv i : PTm n ∈ PUniv j
+| T_Univ n Γ i :
+  ⊢ Γ ->
+  Γ ⊢ PUniv i : PTm n ∈ PUniv (S i)
 
-| T_Conv n Γ (a : PTm n) A B i :
+| T_Conv n Γ (a : PTm n) A B :
   Γ ⊢ a ∈ A ->
-  Γ ⊢ A ≡ B ∈ PUniv i ->
+  Γ ⊢ A ≲ B ->
   Γ ⊢ a ∈ B
 
 with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
@@ -92,20 +93,38 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ a ≡ b ∈ PBind PSig A B ->
   Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B
 
-| E_Conv n Γ (a b : PTm n) A B i :
+| E_Conv n Γ (a b : PTm n) A B :
   Γ ⊢ a ≡ b ∈ A ->
-  Γ ⊢ B ∈ PUniv i ->
-  Γ ⊢ A ≡ B ∈ PUniv i ->
+  Γ ⊢ A ≲ B ->
   Γ ⊢ a ≡ b ∈ B
 
-| E_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
-  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
-  Γ ⊢ A0 ≡ A1 ∈ PUniv i
+with LEq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
+| Su_Transitive  n Γ (A B C : PTm n) :
+  Γ ⊢ A ≲ B ->
+  Γ ⊢ B ≲ C ->
+  Γ ⊢ A ≲ C
 
-| E_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
-  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
+| Su_Univ n Γ i j :
+  i <= j ->
+  Γ ⊢ PUniv i : PTm n ≲ PUniv j
+
+| Su_Bind n Γ p (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ A1 ≲ A0 ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≲ B1 ->
+  Γ ⊢ PBind p A0 B0 ≲ PBind p A1 B1
+
+| Su_Eq n Γ (A : PTm n) B i  :
+  Γ ⊢ A ≡ B ∈ PUniv i ->
+  Γ ⊢ A ≲ B
+
+| Su_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ PBind p A0 B0 ≲ PBind p A1 B1 ->
+  Γ ⊢ A1 ≲ A0
+
+| Su_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ PBind p A0 B0 ≲ PBind p A1 B1 ->
   Γ ⊢ a0 ≡ a1 ∈ A0 ->
-  Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≡ subst_PTm (scons a1 VarPTm) B1 ∈ PUniv i
+  Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≲ subst_PTm (scons a1 VarPTm) B1
 
 with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
 | Wff_Nil :
@@ -117,10 +136,17 @@ with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
   ⊢ funcomp (ren_PTm shift) (scons A Γ)
 
 where
-"Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A).
+"Γ ⊢ a ∈ A" := (Wt Γ a A) and "⊢ Γ" := (Wff Γ) and "Γ ⊢ a ≡ b ∈ A" := (Eq Γ a b A) and "Γ ⊢ A ≲ B" := (LEq Γ A B).
 
 Scheme wf_ind := Induction for Wff Sort Prop
   with wt_ind := Induction for Wt Sort Prop
-  with eq_ind := Induction for Eq Sort Prop.
+  with eq_ind := Induction for Eq Sort Prop
+  with le_ind := Induction for LEq Sort Prop.
 
-Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind.
+Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind, le_ind.
+
+(* Lemma lem : *)
+(*   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> ...) /\ *)
+(*   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> ...)  /\ *)
+(*   (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> ...). *)
+(* Proof. apply wt_mutual. ... *)

From cf2726be8d163ccf689c307751ce0db2425aea88 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Feb 2025 23:41:38 -0500
Subject: [PATCH 052/210] Add subtyping

---
 theories/fp_red.v | 156 ++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 156 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 9afad56..3a4981d 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1774,6 +1774,17 @@ Module REReds.
       eauto using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl.
   Qed.
 
+  Lemma cong' n m (a b : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
+    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).
+  Proof.
+    move => h0 h1.
+    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.
+
 End REReds.
 
 Module LoRed.
@@ -2286,3 +2297,148 @@ Module DJoin.
   Qed.
 
 End DJoin.
+
+
+Module Sub1.
+  Inductive R {n} : PTm n -> PTm n -> Prop :=
+  | Refl a :
+    R a a
+  | Univ i j :
+    i <= j ->
+    R (PUniv i) (PUniv j)
+  | BindCong p A0 A1 B0 B1 :
+    R A1 A0 ->
+    R B0 B1 ->
+    R (PBind p A0 B0) (PBind p A1 B1).
+
+  Lemma transitive0 n (A B C : PTm n) :
+    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) :
+    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.
+  Proof. sfirstorder. Qed.
+
+  Lemma commutativity0 n (A B C : PTm n) :
+    R A B ->
+    (RERed.R A C ->
+    exists D, RERed.R B D /\ R C D) /\
+    (RERed.R B C ->
+    exists D, RERed.R A D /\ R D C).
+  Proof.
+    move => h. move : C.
+    elim : n A B / h.
+    - sfirstorder.
+    - hauto lq:on inv:RERed.R.
+    - hauto lq:on ctrs:RERed.R, R inv:RERed.R.
+  Qed.
+
+  Lemma commutativity_Ls n (A B C : PTm n) :
+    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.
+  Qed.
+
+  Lemma commutativity_Rs n (A B C : PTm n) :
+    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.
+  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).
+  Proof.
+    apply sn_mutual; hauto lq:on inv:R ctrs:SN.
+  Qed.
+
+End Sub1.
+
+(* Module Sub01. *)
+(*   Definition R {n} (a b : PTm n) := a = b \/ Sub1.R a b. *)
+
+(*   Lemma refl n (a : PTm n) : R a a. *)
+(*   Proof. sfirstorder. Qed. *)
+
+(*   Lemma sn_preservation0 : forall n (a b : PTm n), R a b -> SN a <-> SN b. *)
+(*   Proof. hauto lq:on use:Sub1.sn_preservation. Qed. *)
+
+(*   Lemma commutativity_Ls n (A B C : PTm n) : *)
+(*     R A B -> *)
+(*     rtc RERed.R A C -> *)
+(*     exists D, rtc RERed.R B D /\ R C D. *)
+(*   Proof. hauto q:on use:Sub1.commutativity_Ls. Qed. *)
+
+(*   Lemma commutativity_Rs n (A B C : PTm n) : *)
+(*     R A B -> *)
+(*     rtc RERed.R B C -> *)
+(*     exists D, rtc RERed.R A D /\ R D C. *)
+(*   Proof. hauto q:on use:Sub1.commutativity_Rs. Qed. *)
+
+(*   Lemma transitive0 n (A B C : PTm n) : *)
+(*     R A B -> (R B C -> R A C) /\ (R C A -> R C B). *)
+(*   Proof. hauto q:on use:Sub1.transitive. Qed. *)
+
+(*   Lemma transitive n (A B C : PTm n) : *)
+(*     R A B -> R B C -> R A C. *)
+(*   Proof. hauto q:on use:transitive0. Qed. *)
+
+(*   Lemma BindCong n p (A0 A1 : PTm n) B0 B1 : *)
+(*     R A1 A0 -> *)
+(*     R B0 B1 -> *)
+(*     R (PBind p A0 B0) (PBind p A1 B1). *)
+(*   Proof. *)
+(*     best use: *)
+
+(* End Sub01. *)
+
+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.
+
+  Lemma refl n (a : PTm n) : R a a.
+  Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
+
+  Lemma transitive n (a b c : PTm n) : 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.
+    move : hb hb'.
+    move : rered_confluence h. repeat move/[apply].
+    move => [b'][hb0]hb1.
+    have [a' ?] : exists a', rtc RERed.R a0 a' /\ Sub1.R a' b' by hauto l:on use:Sub1.commutativity_Rs.
+    have [c' ?] : exists a', rtc RERed.R c0 a' /\ Sub1.R b' a' by hauto l:on use:Sub1.commutativity_Ls.
+    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.
+  Proof. sfirstorder. Qed.
+
+  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+    R A1 A0 ->
+    R B0 B1 ->
+    R (PBind p A0 B0) (PBind p A1 B1).
+  Proof.
+    rewrite /R.
+    move => [A][A'][h0][h1]h2.
+    move => [B][B'][h3][h4]h5.
+    exists (PBind p A' B), (PBind p A B').
+    repeat split; eauto using REReds.BindCong, Sub1.BindCong.
+  Qed.
+
+  Lemma UnivCong n i j :
+    i <= j ->
+    @R n (PUniv i) (PUniv j).
+  Proof. hauto lq:on ctrs:Sub1.R, rtc. Qed.
+
+End Sub.

From 2f4ea2c78b0b27dbdebdcc656fb1aa4f5cb54524 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 7 Feb 2025 00:39:34 -0500
Subject: [PATCH 053/210] Add more noconfusion lemmas for untyped subtyping

---
 theories/fp_red.v | 109 ++++++++++++++++++++++-------------
 theories/logrel.v | 142 +++++++++++++++++++++++++++++-----------------
 2 files changed, 162 insertions(+), 89 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3a4981d..3c36607 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2363,46 +2363,20 @@ Module Sub1.
     apply sn_mutual; hauto lq:on inv:R ctrs:SN.
   Qed.
 
+  Lemma bind_preservation n (a b : PTm n) :
+    R a b -> isbind a = isbind b.
+  Proof. hauto q:on inv:R. Qed.
+
+  Lemma univ_preservation n (a b : PTm n) :
+    R a b -> isuniv a = isuniv b.
+  Proof. hauto q:on inv:R. Qed.
+
+  Lemma sne_preservation n (a b : PTm n) :
+    R a b -> SNe a <-> SNe b.
+  Proof. hfcrush use:sn_preservation. Qed.
+
 End Sub1.
 
-(* Module Sub01. *)
-(*   Definition R {n} (a b : PTm n) := a = b \/ Sub1.R a b. *)
-
-(*   Lemma refl n (a : PTm n) : R a a. *)
-(*   Proof. sfirstorder. Qed. *)
-
-(*   Lemma sn_preservation0 : forall n (a b : PTm n), R a b -> SN a <-> SN b. *)
-(*   Proof. hauto lq:on use:Sub1.sn_preservation. Qed. *)
-
-(*   Lemma commutativity_Ls n (A B C : PTm n) : *)
-(*     R A B -> *)
-(*     rtc RERed.R A C -> *)
-(*     exists D, rtc RERed.R B D /\ R C D. *)
-(*   Proof. hauto q:on use:Sub1.commutativity_Ls. Qed. *)
-
-(*   Lemma commutativity_Rs n (A B C : PTm n) : *)
-(*     R A B -> *)
-(*     rtc RERed.R B C -> *)
-(*     exists D, rtc RERed.R A D /\ R D C. *)
-(*   Proof. hauto q:on use:Sub1.commutativity_Rs. Qed. *)
-
-(*   Lemma transitive0 n (A B C : PTm n) : *)
-(*     R A B -> (R B C -> R A C) /\ (R C A -> R C B). *)
-(*   Proof. hauto q:on use:Sub1.transitive. Qed. *)
-
-(*   Lemma transitive n (A B C : PTm n) : *)
-(*     R A B -> R B C -> R A C. *)
-(*   Proof. hauto q:on use:transitive0. Qed. *)
-
-(*   Lemma BindCong n p (A0 A1 : PTm n) B0 B1 : *)
-(*     R A1 A0 -> *)
-(*     R B0 B1 -> *)
-(*     R (PBind p A0 B0) (PBind p A1 B1). *)
-(*   Proof. *)
-(*     best use: *)
-
-(* End Sub01. *)
-
 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.
 
@@ -2441,4 +2415,63 @@ Module Sub.
     @R n (PUniv i) (PUniv j).
   Proof. hauto lq:on ctrs:Sub1.R, rtc. Qed.
 
+  Lemma sne_bind_noconf n (a b : PTm n) :
+    R a b -> SNe a -> isbind b -> False.
+  Proof.
+    rewrite /R.
+    move => [c[d] [? []]] *.
+    have : SNe c /\ isbind c by
+      hauto l:on use:REReds.sne_preservation, REReds.bind_preservation, Sub1.sne_preservation, Sub1.bind_preservation.
+    qauto l:on inv:SNe.
+  Qed.
+
+  Lemma bind_sne_noconf n (a b : PTm n) :
+    R a b -> SNe b -> isbind a -> False.
+  Proof.
+    rewrite /R.
+    move => [c[d] [? []]] *.
+    have : SNe c /\ isbind c by
+      hauto l:on use:REReds.sne_preservation, REReds.bind_preservation, Sub1.sne_preservation, Sub1.bind_preservation.
+    qauto l:on inv:SNe.
+  Qed.
+
+  Lemma sne_univ_noconf n (a b : PTm n) :
+    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) :
+    R a b -> SNe b -> isuniv a -> False.
+  Proof.
+    move => [c[d] [? []]] *.
+    have ? : SNe d by eauto using REReds.sne_preservation.
+    have : SNe c by sfirstorder use:Sub1.sne_preservation.
+    have : isuniv c by sfirstorder use:REReds.univ_preservation.
+    clear. case : c => //=. inversion 2.
+  Qed.
+
+  Lemma bind_univ_noconf n (a b : PTm n) :
+    R a b -> isbind a -> isuniv b -> False.
+  Proof.
+    move => [c[d] [h0 [h1 h1']]] h2 h3.
+    have : isbind c /\ isuniv c by
+      hauto l:on use:REReds.bind_preservation,
+            REReds.univ_preservation, Sub1.bind_preservation, Sub1.univ_preservation.
+    move : h2 h3. clear.
+    case : c => //=; itauto.
+  Qed.
+
+  Lemma univ_bind_noconf n (a b : PTm n) :
+    R a b -> isbind a -> isuniv b -> False.
+  Proof.
+    move => [c[d] [h0 [h1 h1']]] h2 h3.
+    have : isbind c /\ isuniv c by
+      hauto l:on use:REReds.bind_preservation,
+            REReds.univ_preservation, Sub1.bind_preservation, Sub1.univ_preservation.
+    move : h2 h3. clear.
+    case : c => //=; itauto.
+  Qed.
+
 End Sub.
diff --git a/theories/logrel.v b/theories/logrel.v
index 4374527..05a5d12 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -261,7 +261,8 @@ Qed.
 Lemma redsns_preservation : forall n a b, @SN n a -> rtc TRedSN a b -> SN b.
 Proof. induction 2; sfirstorder use:redsn_preservation_mutual ctrs:rtc. Qed.
 
-#[export]Hint Resolve DJoin.sne_bind_noconf DJoin.sne_univ_noconf DJoin.bind_univ_noconf : noconf.
+#[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: noconf.
 
 Lemma InterpUniv_SNe_inv n i (A : PTm n) PA :
   SNe A ->
@@ -316,68 +317,107 @@ Proof.
     hauto lq:on rew:off.
 Qed.
 
+Lemma InterpUniv_Sub' n i (A B : PTm n) PA PB :
+  ⟦ A ⟧ i ↘ PA ->
+  ⟦ B ⟧ i ↘ PB ->
+  ((Sub.R A B -> forall x, PA x -> PB x) /\
+  (Sub.R B A -> forall x, PB x -> PA x)).
+Proof.
+  move => hA.
+  move : i A PA hA B PB.
+  apply : InterpUniv_ind.
+  - move => i A hA B PB hPB. split.
+    + move => hAB a ha.
+      have [? ?] : SN B /\ SN A by hauto l:on use:adequacy.
+      move /InterpUniv_case : hPB.
+      move => [H [/DJoin.FromRedSNs h [h1 h0]]].
+      have {h}{}hAB : Sub.R A H by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
+      have {}h0 : SNe H.
+      suff : ~ isbind H /\ ~ isuniv H by itauto.
+      move : hA hAB. clear. hauto lq:on db:noconf.
+      move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
+      tauto.
+    + move => hAB a ha.
+      have [? ?] : SN B /\ SN A by hauto l:on use:adequacy.
+      move /InterpUniv_case : hPB.
+      move => [H [/DJoin.FromRedSNs h [h1 h0]]].
+      have {h}{}hAB : Sub.R H A by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
+      have {}h0 : SNe H.
+      suff : ~ isbind H /\ ~ isuniv H by itauto.
+      move : hAB hA h0. clear. hauto lq:on db:noconf.
+      move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
+      tauto.
+  -
+  (* - move => i p A B PA PF hPA ihPA hTot hRes ihPF U PU hU. *)
+  (*   have hU' : SN U by hauto l:on use:adequacy. *)
+  (*   move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU. *)
+  (*   have {hU} {}h : Sub.R (PBind p A B) H \/ Sub.R H (PBind p A B) *)
+  (*     by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive. *)
+  (*   have{h0} : isbind H. *)
+  (*   suff : ~ SNe H /\ ~ isuniv H by itauto. *)
+  (*   have : isbind (PBind p A B) by scongruence. *)
+  (*   (* hauto l:on use: DJoin.sne_bind_noconf, DJoin.bind_univ_noconf, DJoin.symmetric. *) *)
+  (*   admit. *)
+  (*   case : H h1 h => //=. *)
+  (*   move => p0 A0 B0 h0. /DJoin.bind_inj. *)
+  (*   move => [? [hA hB]] _. subst. *)
+  (*   move /InterpUniv_Bind_inv : h0. *)
+  (*   move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst. *)
+  (*   have ? : PA0 = PA by qauto l:on. subst. *)
+  (*   have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'. *)
+  (*   move => a PB PB' ha hPB hPB'. apply : ihPF; eauto. *)
+  (*   by apply DJoin.substing. *)
+  (*   move => h. extensionality b. apply propositional_extensionality. *)
+  (*   hauto l:on use:bindspace_iff. *)
+  (* - move => i j jlti ih B PB hPB. *)
+  (*   have ? : SN B by hauto l:on use:adequacy. *)
+  (*   move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]]. *)
+  (*   move => hj. *)
+  (*   have {hj}{}h : DJoin.R (PUniv j) H by eauto using DJoin.transitive. *)
+  (*   have {h0} : isuniv H. *)
+  (*   suff : ~ SNe H /\ ~ isbind H by tauto. *)
+  (*   hauto l:on use: DJoin.sne_univ_noconf, DJoin.bind_univ_noconf, DJoin.symmetric. *)
+  (*   case : H h1 h => //=. *)
+  (*   move => j' hPB h _. *)
+  (*   have {}h : j' = j by hauto lq:on use: DJoin.univ_inj. subst. *)
+  (*   hauto lq:on use:InterpUniv_Univ_inv. *)
+  (* - move => i A A0 PA hr hPA ihPA B PB hPB hAB. *)
+  (*   have /DJoin.symmetric ? : DJoin.R A A0 by hauto lq:on rew:off ctrs:rtc use:DJoin.FromRedSNs. *)
+  (*   have ? : SN A0 by hauto l:on use:adequacy. *)
+  (*   have ? : SN A by eauto using N_Exp. *)
+  (*   have : DJoin.R A0 B by eauto using DJoin.transitive. *)
+  (*   eauto. *)
+(* Qed. *)
+Admitted.
+
+Lemma InterpUniv_Sub n i (A B : PTm n) PA PB :
+  ⟦ A ⟧ i ↘ PA ->
+  ⟦ B ⟧ i ↘ PB ->
+  Sub.R A B -> forall x, PA x -> PB x.
+Proof.
+  move : InterpUniv_Sub'. repeat move/[apply].
+  move => [+ _]. apply.
+Qed.
+
 Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
   DJoin.R A B ->
   PA = PB.
 Proof.
-  move => hA.
-  move : i A PA hA B PB.
-  apply : InterpUniv_ind.
-  - move => i A hA B PB hPB hAB.
-    have [*] : SN B /\ SN A by hauto l:on use:adequacy.
-    move /InterpUniv_case : hPB.
-    move => [H [/DJoin.FromRedSNs h [h1 h0]]].
-    have {hAB} {}h : DJoin.R A H by eauto using DJoin.transitive.
-    have {}h0 : SNe H.
-    suff : ~ isbind H /\ ~ isuniv H by itauto.
-    move : h hA. clear. hauto lq:on db:noconf.
-    hauto lq:on use:InterpUniv_SNe_inv.
-  - move => i p A B PA PF hPA ihPA hTot hRes ihPF U PU hU.
-    have hU' : SN U by hauto l:on use:adequacy.
-    move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU.
-    have {hU} {}h : DJoin.R (PBind p A B) H by eauto using DJoin.transitive.
-    have{h0} : isbind H.
-    suff : ~ SNe H /\ ~ isuniv H by itauto.
-    have : isbind (PBind p A B) by scongruence.
-    hauto l:on use: DJoin.sne_bind_noconf, DJoin.bind_univ_noconf, DJoin.symmetric.
-    case : H h1 h => //=.
-    move => p0 A0 B0 h0 /DJoin.bind_inj.
-    move => [? [hA hB]] _. subst.
-    move /InterpUniv_Bind_inv : h0.
-    move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst.
-    have ? : PA0 = PA by qauto l:on. subst.
-    have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'.
-    move => a PB PB' ha hPB hPB'. apply : ihPF; eauto.
-    by apply DJoin.substing.
-    move => h. extensionality b. apply propositional_extensionality.
-    hauto l:on use:bindspace_iff.
-  - move => i j jlti ih B PB hPB.
-    have ? : SN B by hauto l:on use:adequacy.
-    move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].
-    move => hj.
-    have {hj}{}h : DJoin.R (PUniv j) H by eauto using DJoin.transitive.
-    have {h0} : isuniv H.
-    suff : ~ SNe H /\ ~ isbind H by tauto.
-    hauto l:on use: DJoin.sne_univ_noconf, DJoin.bind_univ_noconf, DJoin.symmetric.
-    case : H h1 h => //=.
-    move => j' hPB h _.
-    have {}h : j' = j by hauto lq:on use: DJoin.univ_inj. subst.
-    hauto lq:on use:InterpUniv_Univ_inv.
-  - move => i A A0 PA hr hPA ihPA B PB hPB hAB.
-    have /DJoin.symmetric ? : DJoin.R A A0 by hauto lq:on rew:off ctrs:rtc use:DJoin.FromRedSNs.
-    have ? : SN A0 by hauto l:on use:adequacy.
-    have ? : SN A by eauto using N_Exp.
-    have : DJoin.R A0 B by eauto using DJoin.transitive.
-    eauto.
+  move => + + /[dup] /Sub.FromJoin + /DJoin.symmetric /Sub.FromJoin.
+  move : InterpUniv_Sub'; repeat move/[apply]. move => h.
+  move => h1 h2.
+  extensionality x.
+  apply propositional_extensionality.
+  firstorder.
 Qed.
 
 Lemma InterpUniv_Functional n i  (A : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ A ⟧ i ↘ PB ->
   PA = PB.
-Proof. hauto use:InterpUniv_Join, DJoin.refl. Qed.
+Proof. hauto l:on use:InterpUniv_Join, DJoin.refl. Qed.
 
 Lemma InterpUniv_Join' n i j (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->

From 707483d4014e9baf1362a195b0059763864dfb96 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 7 Feb 2025 00:43:12 -0500
Subject: [PATCH 054/210] Add injectivity for subtyping

---
 theories/fp_red.v | 11 +++++++++++
 1 file changed, 11 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3c36607..cffde37 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2474,4 +2474,15 @@ Module Sub.
     case : c => //=; itauto.
   Qed.
 
+  Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
+    R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
+    p0 = p1 /\ R A1 A0 /\ R B0 B1.
+  Proof.
+    rewrite /R.
+    move => [u][v][h0][h1]h.
+    move /REReds.bind_inv : h0 => [A2][B2][?][h00]h01. subst.
+    move /REReds.bind_inv : h1 => [A3][B3][?][h10]h11. subst.
+    inversion h; subst; sfirstorder.
+  Qed.
+
 End Sub.

From 4bc08e18977db9452b98ae0d1646d616d5fddd7e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 7 Feb 2025 01:19:19 -0500
Subject: [PATCH 055/210] Add one interpuniv sub case

---
 theories/logrel.v | 43 +++++++++++++++++++++----------------------
 1 file changed, 21 insertions(+), 22 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 05a5d12..62663c7 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -347,28 +347,27 @@ Proof.
       move : hAB hA h0. clear. hauto lq:on db:noconf.
       move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
       tauto.
-  -
-  (* - move => i p A B PA PF hPA ihPA hTot hRes ihPF U PU hU. *)
-  (*   have hU' : SN U by hauto l:on use:adequacy. *)
-  (*   move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU. *)
-  (*   have {hU} {}h : Sub.R (PBind p A B) H \/ Sub.R H (PBind p A B) *)
-  (*     by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive. *)
-  (*   have{h0} : isbind H. *)
-  (*   suff : ~ SNe H /\ ~ isuniv H by itauto. *)
-  (*   have : isbind (PBind p A B) by scongruence. *)
-  (*   (* hauto l:on use: DJoin.sne_bind_noconf, DJoin.bind_univ_noconf, DJoin.symmetric. *) *)
-  (*   admit. *)
-  (*   case : H h1 h => //=. *)
-  (*   move => p0 A0 B0 h0. /DJoin.bind_inj. *)
-  (*   move => [? [hA hB]] _. subst. *)
-  (*   move /InterpUniv_Bind_inv : h0. *)
-  (*   move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst. *)
-  (*   have ? : PA0 = PA by qauto l:on. subst. *)
-  (*   have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'. *)
-  (*   move => a PB PB' ha hPB hPB'. apply : ihPF; eauto. *)
-  (*   by apply DJoin.substing. *)
-  (*   move => h. extensionality b. apply propositional_extensionality. *)
-  (*   hauto l:on use:bindspace_iff. *)
+  - move => i p A B PA PF hPA ihPA hTot hRes ihPF U PU hU. split.
+    + have hU' : SN U by hauto l:on use:adequacy.
+      move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU.
+      have {hU} {}h : Sub.R (PBind p A B) H \/ Sub.R H (PBind p A B)
+        by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
+      have{h0} : isbind H.
+      suff : ~ SNe H /\ ~ isuniv H by itauto.
+      have : isbind (PBind p A B) by scongruence.
+      hauto l:on db:noconf.
+      admit.
+      case : H h1 h => //=.
+      move => p0 A0 B0 h0. /DJoin.bind_inj.
+      move => [? [hA hB]] _. subst.
+      move /InterpUniv_Bind_inv : h0.
+      move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst.
+      have ? : PA0 = PA by qauto l:on. subst.
+      have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'.
+      move => a PB PB' ha hPB hPB'. apply : ihPF; eauto.
+      by apply DJoin.substing.
+      move => h. extensionality b. apply propositional_extensionality.
+      hauto l:on use:bindspace_iff.
   (* - move => i j jlti ih B PB hPB. *)
   (*   have ? : SN B by hauto l:on use:adequacy. *)
   (*   move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]]. *)

From 0746e9a354b679952aaad66b069eac43efbb83e6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 7 Feb 2025 16:45:58 -0500
Subject: [PATCH 056/210] Finish subtyping semantics

---
 theories/fp_red.v |  69 +++++++++++++++++++++-----
 theories/logrel.v | 124 +++++++++++++++++++++++++++++-----------------
 2 files changed, 135 insertions(+), 58 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index cffde37..28b698c 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2306,10 +2306,14 @@ Module Sub1.
   | Univ i j :
     i <= j ->
     R (PUniv i) (PUniv j)
-  | BindCong p A0 A1 B0 B1 :
+  | PiCong A0 A1 B0 B1 :
     R A1 A0 ->
     R B0 B1 ->
-    R (PBind p A0 B0) (PBind p A1 B1).
+    R (PBind PPi A0 B0) (PBind PPi A1 B1)
+  | SigCong A0 A1 B0 B1 :
+    R A0 A1 ->
+    R B0 B1 ->
+    R (PBind PSig A0 B0) (PBind PSig A1 B1).
 
   Lemma transitive0 n (A B C : PTm n) :
     R A B -> (R B C -> R A C) /\ (R C A -> R C B).
@@ -2333,10 +2337,7 @@ Module Sub1.
     exists D, RERed.R A D /\ R D C).
   Proof.
     move => h. move : C.
-    elim : n A B / h.
-    - sfirstorder.
-    - hauto lq:on inv:RERed.R.
-    - hauto lq:on ctrs:RERed.R, R inv:RERed.R.
+    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) :
@@ -2375,6 +2376,17 @@ Module Sub1.
     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).
+  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).
+  Proof. move => h. move : m ρ. elim : n a b /h; hauto lq:on ctrs:R. Qed.
+
 End Sub1.
 
 Module Sub.
@@ -2398,16 +2410,28 @@ Module Sub.
   Lemma FromJoin n (a b : PTm n) : DJoin.R a b -> R a b.
   Proof. sfirstorder. Qed.
 
-  Lemma BindCong n p (A0 A1 : PTm n) B0 B1 :
+  Lemma PiCong n (A0 A1 : PTm n) B0 B1 :
     R A1 A0 ->
     R B0 B1 ->
-    R (PBind p A0 B0) (PBind p A1 B1).
+    R (PBind PPi A0 B0) (PBind PPi A1 B1).
   Proof.
     rewrite /R.
     move => [A][A'][h0][h1]h2.
     move => [B][B'][h3][h4]h5.
-    exists (PBind p A' B), (PBind p A B').
-    repeat split; eauto using REReds.BindCong, Sub1.BindCong.
+    exists (PBind PPi A' B), (PBind PPi A B').
+    repeat split; eauto using REReds.BindCong, Sub1.PiCong.
+  Qed.
+
+  Lemma SigCong n (A0 A1 : PTm n) B0 B1 :
+    R A0 A1 ->
+    R B0 B1 ->
+    R (PBind PSig A0 B0) (PBind PSig A1 B1).
+  Proof.
+    rewrite /R.
+    move => [A][A'][h0][h1]h2.
+    move => [B][B'][h3][h4]h5.
+    exists (PBind PSig A B), (PBind PSig A' B').
+    repeat split; eauto using REReds.BindCong, Sub1.SigCong.
   Qed.
 
   Lemma UnivCong n i j :
@@ -2464,7 +2488,7 @@ Module Sub.
   Qed.
 
   Lemma univ_bind_noconf n (a b : PTm n) :
-    R a b -> isbind a -> isuniv b -> False.
+    R a b -> isbind b -> isuniv a -> False.
   Proof.
     move => [c[d] [h0 [h1 h1']]] h2 h3.
     have : isbind c /\ isuniv c by
@@ -2476,13 +2500,32 @@ Module Sub.
 
   Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
     R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
-    p0 = p1 /\ R A1 A0 /\ R B0 B1.
+    p0 = p1 /\ (if p0 is PPi then R A1 A0 else R A0 A1) /\ R B0 B1.
   Proof.
     rewrite /R.
     move => [u][v][h0][h1]h.
     move /REReds.bind_inv : h0 => [A2][B2][?][h00]h01. subst.
     move /REReds.bind_inv : h1 => [A3][B3][?][h10]h11. subst.
-    inversion h; subst; sfirstorder.
+    inversion h; subst; sauto lq:on.
+  Qed.
+
+  Lemma univ_inj n i j :
+    @R n (PUniv i) (PUniv j)  -> i <= j.
+  Proof.
+    sauto lq:on rew:off use:REReds.univ_inv.
+  Qed.
+
+  Lemma cong n m (a b : PTm (S n)) c d (ρ : fin n -> PTm m) :
+    R a b -> DJoin.R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) b).
+  Proof.
+    rewrite /R.
+    move => [a0][b0][h0][h1]h2.
+    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.
+    eauto using Sub1.substing.
   Qed.
 
 End Sub.
diff --git a/theories/logrel.v b/theories/logrel.v
index 62663c7..ec76c75 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -40,7 +40,6 @@ 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) :
   PF = I j ->
   j < i ->
@@ -295,26 +294,24 @@ Proof.
   subst. hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma bindspace_iff n p (PA : PTm n -> Prop) PF PF0 b  :
-  (forall (a : PTm n) (PB PB0 : PTm n -> Prop), PA a -> PF a PB -> PF0 a PB0 -> PB = PB0) ->
+Lemma bindspace_impl n p (PA PA0 : PTm n -> 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, PA a -> exists PB, PF a PB) ->
-  (forall a, PA a -> exists PB0, PF0 a PB0) ->
-  (BindSpace p PA PF b <-> BindSpace p PA PF0 b).
+  (forall a, PA0 a -> exists PB0, PF0 a PB0) ->
+  (BindSpace p PA PF b -> BindSpace p PA0 PF0 b).
 Proof.
-  rewrite /BindSpace => h hPF hPF0.
-  case : p => /=.
+  rewrite /BindSpace => hSA h hPF hPF0.
+  case : p hSA => /= hSA.
   - rewrite /ProdSpace.
-    split.
     move => h1 a PB ha hPF'.
-    specialize hPF with (1 := ha).
+    have {}/hPF : PA a by sfirstorder.
     specialize hPF0 with (1 := ha).
-    sblast.
-    move => ? a PB ha.
-    specialize hPF with (1 := ha).
-    specialize hPF0 with (1 := ha).
-    sblast.
+    hauto lq:on.
   - rewrite /SumSpace.
-    hauto lq:on rew:off.
+    case. sfirstorder.
+    move => [a0][b0][h0][h1]h2. right.
+    hauto lq:on.
 Qed.
 
 Lemma InterpUniv_Sub' n i (A B : PTm n) PA PB :
@@ -350,44 +347,76 @@ Proof.
   - move => i p A B PA PF hPA ihPA hTot hRes ihPF U PU hU. split.
     + have hU' : SN U by hauto l:on use:adequacy.
       move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU.
-      have {hU} {}h : Sub.R (PBind p A B) H \/ Sub.R H (PBind p A B)
+      have {hU} {}h : Sub.R (PBind p A B) H
         by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have{h0} : isbind H.
       suff : ~ SNe H /\ ~ isuniv H by itauto.
       have : isbind (PBind p A B) by scongruence.
-      hauto l:on db:noconf.
-      admit.
+      move : h. clear. hauto l:on db:noconf.
       case : H h1 h => //=.
-      move => p0 A0 B0 h0. /DJoin.bind_inj.
+      move => p0 A0 B0 h0 /Sub.bind_inj.
       move => [? [hA hB]] _. subst.
       move /InterpUniv_Bind_inv : h0.
       move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst.
-      have ? : PA0 = PA by qauto l:on. subst.
-      have : forall a PB PB', PA a -> PF a PB -> PF0 a PB' -> PB = PB'.
-      move => a PB PB' ha hPB hPB'. apply : ihPF; eauto.
-      by apply DJoin.substing.
-      move => h. extensionality b. apply propositional_extensionality.
-      hauto l:on use:bindspace_iff.
-  (* - move => i j jlti ih B PB hPB. *)
-  (*   have ? : SN B by hauto l:on use:adequacy. *)
-  (*   move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]]. *)
-  (*   move => hj. *)
-  (*   have {hj}{}h : DJoin.R (PUniv j) H by eauto using DJoin.transitive. *)
-  (*   have {h0} : isuniv H. *)
-  (*   suff : ~ SNe H /\ ~ isbind H by tauto. *)
-  (*   hauto l:on use: DJoin.sne_univ_noconf, DJoin.bind_univ_noconf, DJoin.symmetric. *)
-  (*   case : H h1 h => //=. *)
-  (*   move => j' hPB h _. *)
-  (*   have {}h : j' = j by hauto lq:on use: DJoin.univ_inj. subst. *)
-  (*   hauto lq:on use:InterpUniv_Univ_inv. *)
-  (* - move => i A A0 PA hr hPA ihPA B PB hPB hAB. *)
-  (*   have /DJoin.symmetric ? : DJoin.R A A0 by hauto lq:on rew:off ctrs:rtc use:DJoin.FromRedSNs. *)
-  (*   have ? : SN A0 by hauto l:on use:adequacy. *)
-  (*   have ? : SN A by eauto using N_Exp. *)
-  (*   have : DJoin.R A0 B by eauto using DJoin.transitive. *)
-  (*   eauto. *)
-(* Qed. *)
-Admitted.
+      move => x. apply bindspace_impl; eauto;[idtac|idtac]. hauto l:on.
+      move => a PB PB' ha hPB hPB'.
+      move : hRes0 hPB'. repeat move/[apply].
+      move : ihPF hPB. repeat move/[apply].
+      move => h. eapply h.
+      apply Sub.cong => //=; eauto using DJoin.refl.
+    + have hU' : SN U by hauto l:on use:adequacy.
+      move /InterpUniv_case : hU => [H [/DJoin.FromRedSNs h [h1 h0]]] hU.
+      have {hU} {}h : Sub.R H (PBind p A B)
+        by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
+      have{h0} : isbind H.
+      suff : ~ SNe H /\ ~ isuniv H by itauto.
+      have : isbind (PBind p A B) by scongruence.
+      move : h. clear. move : (PBind p A B). hauto lq:on db:noconf.
+      case : H h1 h => //=.
+      move => p0 A0 B0 h0 /Sub.bind_inj.
+      move => [? [hA hB]] _. subst.
+      move /InterpUniv_Bind_inv : h0.
+      move => [PA0][PF0][hPA0][hTot0][hRes0 ?]. subst.
+      move => x. apply bindspace_impl; eauto;[idtac|idtac]. hauto l:on.
+      move => a PB PB' ha hPB hPB'.
+      eapply ihPF; eauto.
+      apply Sub.cong => //=; eauto using DJoin.refl.
+  - move => i j jlti ih B PB hPB. split.
+    + have ? : SN B by hauto l:on use:adequacy.
+      move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].
+      move => hj.
+      have {hj}{}h : Sub.R (PUniv j) H by eauto using Sub.transitive, Sub.FromJoin.
+      have {h0} : isuniv H.
+      suff : ~ SNe H /\ ~ isbind H by tauto. move : h. clear. hauto lq:on db:noconf.
+      case : H h1 h => //=.
+      move => j' hPB h _.
+      have {}h : j <= j' by hauto lq:on use: Sub.univ_inj. subst.
+      move /InterpUniv_Univ_inv : hPB => [? ?]. subst.
+      have ? : j <= i by lia.
+      move => A. hauto l:on use:InterpUniv_cumulative.
+    + have ? : SN B by hauto l:on use:adequacy.
+      move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].
+      move => hj.
+      have {hj}{}h : Sub.R H (PUniv j) by eauto using Sub.transitive, Sub.FromJoin, DJoin.symmetric.
+      have {h0} : isuniv H.
+      suff : ~ SNe H /\ ~ isbind H by tauto. move : h. clear. hauto lq:on db:noconf.
+      case : H h1 h => //=.
+      move => j' hPB h _.
+      have {}h : j' <= j by hauto lq:on use: Sub.univ_inj.
+      move /InterpUniv_Univ_inv : hPB => [? ?]. subst.
+      move => A. hauto l:on use:InterpUniv_cumulative.
+  - move => i A A0 PA hr hPA ihPA B PB hPB.
+    have ? : SN A by sauto lq:on use:adequacy.
+    split.
+    + move => ?.
+      have {}hr : Sub.R A0 A by hauto lq:on  ctrs:rtc use:DJoin.FromRedSNs, DJoin.symmetric, Sub.FromJoin.
+      have : Sub.R A0 B by eauto using Sub.transitive.
+      qauto l:on.
+    + move => ?.
+      have {}hr : Sub.R A A0 by hauto lq:on  ctrs:rtc use:DJoin.FromRedSNs, DJoin.symmetric, Sub.FromJoin.
+      have : Sub.R B A0 by eauto using Sub.transitive.
+      qauto l:on.
+Qed.
 
 Lemma InterpUniv_Sub n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
@@ -478,6 +507,9 @@ 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).
 Notation "Γ ⊨ a ≡ b ∈ A" := (SemEq Γ a b A) (at level 70).
 
+Definition SemLEq {n} Γ (a b A : PTm n) := Sub.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" := (SemLEq Γ a b A) (at level 70).
+
 Lemma SemEq_SemWt n Γ (a b A : PTm n) : Γ ⊨ a ≡ b ∈ A -> Γ ⊨ a ∈ A /\ Γ ⊨ b ∈ A /\ DJoin.R a b.
 Proof. hauto lq:on rew:off unfold:SemEq, SemWt. Qed.
 
@@ -492,6 +524,8 @@ Proof.
   hauto lq:on.
 Qed.
 
+Lemma SemEq_SemLEq n Γ
+
 (* Semantic context wellformedness *)
 Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ ⊨ Γ i ∈ PUniv j.
 Notation "⊨ Γ" := (SemWff Γ) (at level 70).

From f483d63f01a22cdf995960baf2403cd71a9159c9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 20:37:46 -0500
Subject: [PATCH 057/210] Fix the definition of semleq

---
 theories/logrel.v | 83 ++++++++++++++++++++++++++++++++++++-----------
 1 file changed, 64 insertions(+), 19 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index ec76c75..084663f 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -418,7 +418,7 @@ Proof.
       qauto l:on.
 Qed.
 
-Lemma InterpUniv_Sub n i (A B : PTm n) PA PB :
+Lemma InterpUniv_Sub0 n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
   Sub.R A B -> forall x, PA x -> PB x.
@@ -427,6 +427,19 @@ Proof.
   move => [+ _]. apply.
 Qed.
 
+Lemma InterpUniv_Sub n i j (A B : PTm n) PA PB :
+  ⟦ A ⟧ i ↘ PA ->
+  ⟦ B ⟧ j ↘ PB ->
+  Sub.R A B -> forall x, PA x -> PB x.
+Proof.
+  have [? ?] : i <= max i j /\ j <= max i j by lia.
+  move => hPA hPB.
+  have : ⟦ B ⟧ (max i j) ↘ PB by eauto using InterpUniv_cumulative.
+  have : ⟦ A ⟧ (max i j) ↘ PA by eauto using InterpUniv_cumulative.
+  move : InterpUniv_Sub0. repeat move/[apply].
+  apply.
+Qed.
+
 Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
   ⟦ A ⟧ i ↘ PA ->
   ⟦ B ⟧ i ↘ PB ->
@@ -507,12 +520,30 @@ 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).
 Notation "Γ ⊨ a ≡ b ∈ A" := (SemEq Γ a b A) (at level 70).
 
-Definition SemLEq {n} Γ (a b A : PTm n) := Sub.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" := (SemLEq Γ 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.
+Notation "Γ ⊨ a ≲ b" := (SemLEq Γ a b) (at level 70).
+
+Lemma SemWt_Univ n Γ (A : PTm n) i  :
+  Γ ⊨ A ∈ PUniv i <->
+  forall ρ, ρ_ok Γ ρ -> exists S, ⟦ subst_PTm ρ A ⟧ i ↘ S.
+Proof.
+  rewrite /SemWt.
+  split.
+  - hauto lq:on rew:off use:InterpUniv_Univ_inv.
+  - move => /[swap] ρ /[apply].
+    move => [PA hPA].
+    exists (S i). eexists.
+    split.
+    + simp InterpUniv. apply InterpExt_Univ. lia.
+    + simpl. eauto.
+Qed.
 
 Lemma SemEq_SemWt n Γ (a b A : PTm n) : Γ ⊨ 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.
+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.
 Proof.
   move => ha hb heq. split => //= ρ hρ.
@@ -524,7 +555,22 @@ Proof.
   hauto lq:on.
 Qed.
 
-Lemma SemEq_SemLEq n Γ
+Lemma SemWt_SemLEq n Γ (A B : PTm n) i j :
+  Γ ⊨ A ∈ PUniv i -> Γ ⊨ B ∈ PUniv j -> Sub.R A B -> Γ ⊨ A ≲ B.
+Proof.
+  move => ha hb heq. split => //.
+  exists (Nat.max i j).
+  have [? ?] : i <= Nat.max i j /\ j <= Nat.max i j by lia.
+  move => ρ hρ.
+  have {}/ha := hρ.
+  have {}/hb := hρ.
+  move => [k][PA][/= /InterpUniv_Univ_inv [? hPA]]hpb.
+  move => [k0][PA0][/= /InterpUniv_Univ_inv [? hPA0]]hpa. subst.
+  move : hpb => [PA]hPA'.
+  move : hpa => [PB]hPB'.
+  exists PB, PA.
+  split; apply : InterpUniv_cumulative; eauto.
+Qed.
 
 (* Semantic context wellformedness *)
 Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ ⊨ Γ i ∈ PUniv j.
@@ -614,21 +660,6 @@ Lemma SemEq_SN_Join n Γ (a b A : PTm n) :
   SN a /\ SN b /\ SN A /\ DJoin.R a b.
 Proof. hauto l:on use:SemEq_SemWt, SemWt_SN. Qed.
 
-Lemma SemWt_Univ n Γ (A : PTm n) i  :
-  Γ ⊨ A ∈ PUniv i <->
-  forall ρ, ρ_ok Γ ρ -> exists S, ⟦ subst_PTm ρ A ⟧ i ↘ S.
-Proof.
-  rewrite /SemWt.
-  split.
-  - hauto lq:on rew:off use:InterpUniv_Univ_inv.
-  - move => /[swap] ρ /[apply].
-    move => [PA hPA].
-    exists (S i). eexists.
-    split.
-    + simp InterpUniv. apply InterpExt_Univ. lia.
-    + simpl. eauto.
-Qed.
-
 (* Structural laws for Semantic context wellformedness *)
 Lemma SemWff_nil : SemWff null.
 Proof. case. Qed.
@@ -861,6 +892,20 @@ Proof.
   hauto l:on use:DJoin.transitive.
 Qed.
 
+Lemma SLEq_Transitive n Γ (A B C : PTm n) :
+  Γ ⊨ A ≲ B ->
+  Γ ⊨ B ≲ C ->
+  Γ ⊨ A ≲ C.
+Proof.
+  move => h0 h1.
+  rewrite /SemLEq in h0 h1.
+  move : h0 => [hAB]h0.
+  move : h1 => [hBC]h1.
+
+  rewrite /SemLEq.
+  split. eauto using Sub.transitive.
+
+
 Lemma Γ_eq_ρ_ok n Γ Δ (ρ : fin n -> PTm 0) : Γ_eq Γ Δ -> ⊨ Γ -> ρ_ok Γ ρ -> ρ_ok Δ ρ.
 Proof.
   move => hΓΔ hΓ h.

From 916e0bcd75b8f6c269c5451205bc4e1d541dd91b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 21:06:51 -0500
Subject: [PATCH 058/210] Change the conversion rules to use Sub instead of Eq

---
 theories/fp_red.v |  7 +++++
 theories/logrel.v | 67 +++++++++++++++++++++++++++++++----------------
 2 files changed, 51 insertions(+), 23 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 28b698c..9998924 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2528,4 +2528,11 @@ 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).
+  Proof.
+    rewrite /R.
+    move => [a0][b0][h0][h1]h2.
+    hauto ctrs:rtc use:REReds.cong', Sub1.substing.
+  Qed.
 End Sub.
diff --git a/theories/logrel.v b/theories/logrel.v
index 084663f..904eb33 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -660,6 +660,11 @@ Lemma SemEq_SN_Join n Γ (a b A : PTm n) :
   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) :
+  Γ ⊨ 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.
@@ -852,23 +857,34 @@ Qed.
 Lemma ST_Conv' n Γ (a : PTm n) A B i :
   Γ ⊨ a ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
-  DJoin.R A B ->
+  Sub.R A B ->
   Γ ⊨ a ∈ B.
 Proof.
   move => ha /SemWt_Univ h h0.
   move => ρ hρ.
-  have {}h0 : DJoin.R (subst_PTm ρ A) (subst_PTm ρ B) by eauto using DJoin.substing.
+  have {}h0 : Sub.R (subst_PTm ρ A) (subst_PTm ρ B) by
+    eauto using Sub.substing.
   move /ha : (hρ){ha} => [m [PA [h1 h2]]].
   move /h : (hρ){h} => [S hS].
-  have ? : PA = S by eauto using InterpUniv_Join'. subst.
-  eauto.
+  have h3 : forall x, PA x -> S x.
+  move : InterpUniv_Sub h0 h1 hS; by repeat move/[apply].
+  hauto lq:on.
 Qed.
 
-Lemma ST_Conv n Γ (a : PTm n) A B i :
+Lemma ST_Conv_E n Γ (a : PTm n) A B i :
   Γ ⊨ a ∈ A ->
-  Γ ⊨ A ≡ B ∈ PUniv i ->
+  Γ ⊨ B ∈ PUniv i ->
+  DJoin.R A B ->
   Γ ⊨ a ∈ B.
-Proof. hauto l:on use:ST_Conv', SemEq_SemWt. Qed.
+Proof.
+  hauto l:on use:ST_Conv', Sub.FromJoin.
+Qed.
+
+Lemma ST_Conv n Γ (a : PTm n) 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 :
   Γ ⊨ a ∈ A ->
@@ -892,19 +908,24 @@ Proof.
   hauto l:on use:DJoin.transitive.
 Qed.
 
+Lemma SLEq_Eq n Γ (A B : PTm n) i :
+  Γ ⊨ A ≡ B ∈ PUniv i ->
+  Γ ⊨ A ≲ B.
+Proof. move /SemEq_SemWt => h.
+       qauto l:on use:SemWt_SemLEq, Sub.FromJoin.
+Qed.
+
 Lemma SLEq_Transitive n Γ (A B C : PTm n) :
   Γ ⊨ A ≲ B ->
   Γ ⊨ B ≲ C ->
   Γ ⊨ A ≲ C.
 Proof.
-  move => h0 h1.
-  rewrite /SemLEq in h0 h1.
-  move : h0 => [hAB]h0.
-  move : h1 => [hBC]h1.
-
-  rewrite /SemLEq.
-  split. eauto using Sub.transitive.
-
+  move => ha hb.
+  apply SemLEq_SemWt in ha, hb.
+  have ? : SN B by hauto l:on use:SemWt_SN.
+  move : ha => [ha0 [i [ha1 ha2]]]. move : hb => [hb0 [j [hb1 hb2]]].
+  qauto l:on use:SemWt_SemLEq, Sub.transitive.
+Qed.
 
 Lemma Γ_eq_ρ_ok n Γ Δ (ρ : fin n -> PTm 0) : Γ_eq Γ Δ -> ⊨ Γ -> ρ_ok Γ ρ -> ρ_ok Δ ρ.
 Proof.
@@ -973,19 +994,19 @@ Qed.
 Lemma SE_Conv' n Γ (a b : PTm n) A B i :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
-  DJoin.R A B ->
+  Sub.R A B ->
   Γ ⊨ a ≡ b ∈ B.
 Proof.
   move /SemEq_SemWt => [ha][hb]he hB hAB.
   apply SemWt_SemEq; eauto using ST_Conv'.
 Qed.
 
-Lemma SE_Conv n Γ (a b : PTm n) A B i :
+Lemma SE_Conv n Γ (a b : PTm n) A B :
   Γ ⊨ a ≡ b ∈ A ->
-  Γ ⊨ A ≡ B ∈ PUniv i ->
+  Γ ⊨ A ≲ B ->
   Γ ⊨ a ≡ b ∈ B.
 Proof.
-  move => h /SemEq_SemWt [h0][h1]he.
+  move => h /SemLEq_SemWt [h0][h1][ha]hb.
   eauto using SE_Conv'.
 Qed.
 
@@ -1016,7 +1037,7 @@ Lemma SE_Pair n Γ (a0 a1 b0 b1 : PTm n) A B i :
   Γ ⊨ PPair a0 b0 ≡ PPair a1 b1 ∈ PBind PSig A B.
 Proof.
   move => h /SemEq_SemWt [ha0][ha1]hae /SemEq_SemWt [hb0][hb1]hbe.
-  apply SemWt_SemEq; eauto using ST_Pair, DJoin.PairCong, SBind_inst, DJoin.cong, ST_Conv', ST_Pair.
+  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 :
@@ -1036,7 +1057,7 @@ Proof.
   move => /SemEq_SemWt [ha][hb]he.
   apply SemWt_SemEq; eauto using DJoin.ProjCong, ST_Proj2.
   have h : Γ ⊨ PProj PR b ∈ subst_PTm (scons (PProj PL b) VarPTm) B by eauto using ST_Proj2.
-  apply : ST_Conv'. apply h.
+  apply : ST_Conv_E. apply h.
   apply : SBind_inst. eauto using ST_Proj1.
   eauto.
   hauto lq:on use: DJoin.cong,  DJoin.ProjCong.
@@ -1051,7 +1072,7 @@ Proof.
   move => hPi.
   move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha.
   apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App.
-  apply : ST_Conv'; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
+  apply : ST_Conv_E; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
 Qed.
 
 Lemma SBind_inv1 n Γ i p (A : PTm n) B :
@@ -1080,7 +1101,7 @@ Proof.
   move : ha => [c [ha ha']].
   apply SemWt_SemEq; eauto using SBind_inst.
   apply SBind_inst with (p := p) (A := A1); eauto.
-  apply : ST_Conv'; eauto. hauto l:on use:SBind_inv1.
+  apply : ST_Conv_E; eauto. hauto l:on use:SBind_inv1.
   rewrite /DJoin.R.
   eexists. split.
   apply : relations.rtc_transitive.

From 02e6c9e025a8b8f4fb1b72e75739a16114811ae1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 22:15:04 -0500
Subject: [PATCH 059/210] Add pi and sig subtyping semantic rules

---
 theories/logrel.v | 169 ++++++++++++++++++++++++++++++++++++----------
 1 file changed, 134 insertions(+), 35 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 904eb33..ed887be 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -512,8 +512,6 @@ Qed.
 Definition ρ_ok {n} (Γ : fin n -> PTm n) (ρ : fin n -> PTm 0) := forall i k PA,
     ⟦ subst_PTm ρ (Γ i) ⟧ k ↘ PA -> PA (ρ i).
 
-Definition Γ_eq {n} (Γ Δ : fin n -> PTm n)  := forall i, DJoin.R (Γ i) (Δ i).
-
 Definition SemWt {n} Γ (a A : PTm n) := forall ρ, ρ_ok Γ ρ -> exists k PA, ⟦ subst_PTm ρ A ⟧ k ↘ PA /\ PA (subst_PTm ρ a).
 Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70).
 
@@ -908,24 +906,7 @@ Proof.
   hauto l:on use:DJoin.transitive.
 Qed.
 
-Lemma SLEq_Eq n Γ (A B : PTm n) i :
-  Γ ⊨ A ≡ B ∈ PUniv i ->
-  Γ ⊨ A ≲ B.
-Proof. move /SemEq_SemWt => h.
-       qauto l:on use:SemWt_SemLEq, Sub.FromJoin.
-Qed.
-
-Lemma SLEq_Transitive n Γ (A B C : PTm n) :
-  Γ ⊨ A ≲ B ->
-  Γ ⊨ B ≲ C ->
-  Γ ⊨ A ≲ C.
-Proof.
-  move => ha hb.
-  apply SemLEq_SemWt in ha, hb.
-  have ? : SN B by hauto l:on use:SemWt_SN.
-  move : ha => [ha0 [i [ha1 ha2]]]. move : hb => [hb0 [j [hb1 hb2]]].
-  qauto l:on use:SemWt_SemLEq, Sub.transitive.
-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.
@@ -942,6 +923,44 @@ Proof.
   hauto l:on use: DJoin.substing.
 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.
@@ -958,7 +977,6 @@ Proof.
   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 Γ)).
@@ -1082,13 +1100,102 @@ Lemma SBind_inv1 n Γ i p (A : PTm n) B :
   hauto lq:on rew:off  use:InterpUniv_Bind_inv.
 Qed.
 
-Lemma SE_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i :
-  Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
-  Γ ⊨ A0 ≡ A1 ∈ PUniv i.
+Lemma SSu_Eq n Γ (A B : PTm n) 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) :
+  Γ ⊨ A ≲ B ->
+  Γ ⊨ B ≲ C ->
+  Γ ⊨ A ≲ C.
 Proof.
-  move /SemEq_SemWt => [h0][h1]he.
-  apply SemWt_SemEq; eauto using SBind_inv1.
-  move /DJoin.bind_inj : he. tauto.
+  move => ha hb.
+  apply SemLEq_SemWt in ha, hb.
+  have ? : SN B by hauto l:on use:SemWt_SN.
+  move : ha => [ha0 [i [ha1 ha2]]]. move : hb => [hb0 [j [hb1 hb2]]].
+  qauto l:on use:SemWt_SemLEq, Sub.transitive.
+Qed.
+
+Lemma ST_Univ n Γ i j :
+  i < j ->
+  Γ ⊨ PUniv i : PTm n ∈ PUniv j.
+Proof.
+  move => ?.
+  apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ.
+Qed.
+
+Lemma SSu_Univ n Γ i j :
+  i <= j ->
+  Γ ⊨ PUniv i : PTm n ≲ 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 :
+  ⊨ Γ ->
+  Γ ⊨ A1 ≲ A0 ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≲ B1 ->
+  Γ ⊨ PBind PPi A0 B0 ≲ PBind PPi A1 B1.
+Proof.
+  move => hΓ 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.
+  move : hA => [hA0][i][hA1]hA2.
+  move : hB => [hB0][j][hB1]hB2.
+  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.
+  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'.
+Qed.
+
+Lemma SSu_Sig n Γ (A0 A1 : PTm n) B0 B1 :
+  ⊨ Γ ->
+  Γ ⊨ A0 ≲ A1 ->
+  funcomp (ren_PTm shift) (scons A1 Γ) ⊨ B0 ≲ B1 ->
+  Γ ⊨ PBind PSig A0 B0 ≲ PBind PSig A1 B1.
+Proof.
+  move => hΓ 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.
+  move : hA => [hA0][i][hA1]hA2.
+  move : hB => [hB0][j][hB1]hB2.
+  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.
+  move => ρ hρ.
+  suff : ρ_ok (funcomp (ren_PTm shift) (scons A1 Γ)) ρ by hauto l:on.
+  apply : Γ_sub_ρ_ok; eauto.
+  hauto lq:on use:Γ_sub_cons'.
+Qed.
+
+Lemma SSu_Pi_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+  Γ ⊨ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
+  Γ ⊨ A1 ≲ A0.
+Proof.
+  move /SemLEq_SemWt => [h0][h1][h2]he.
+  apply : SemWt_SemLEq; eauto using SBind_inv1.
+  hauto lq:on rew:off use:Sub.bind_inj.
+Qed.
+
+Lemma SSu_Sig_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+  Γ ⊨ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
+  Γ ⊨ A0 ≲ A1.
+Proof.
+  move /SemLEq_SemWt => [h0][h1][h2]he.
+  apply : SemWt_SemLEq; eauto using SBind_inv1.
+  hauto lq:on rew:off use:Sub.bind_inj.
 Qed.
 
 Lemma SE_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
@@ -1114,14 +1221,6 @@ Proof.
   hauto lq:on ctrs:rtc inv:option.
 Qed.
 
-Lemma ST_Univ n Γ i j :
-  i < j ->
-  Γ ⊨ PUniv i : PTm n ∈ PUniv j.
-Proof.
-  move => ?.
-  apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ.
-Qed.
-
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
   SE_Conv SE_Bind_Proj1 SE_Bind_Proj2 SemWff_nil SemWff_cons : sem.

From 5996c45526b350e4ea78658fce4d40c8f106d330 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 22:38:28 -0500
Subject: [PATCH 060/210] Finish the semantic projection rules

---
 theories/logrel.v | 45 ++++++++++++++++++++++++---------------------
 1 file changed, 24 insertions(+), 21 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index ed887be..70c0706 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1198,29 +1198,32 @@ Proof.
   hauto lq:on rew:off use:Sub.bind_inj.
 Qed.
 
-Lemma SE_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i :
-  Γ ⊨ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i ->
-  Γ ⊨ a0 ≡ a1 ∈ A0 ->
-  Γ ⊨ subst_PTm (scons a0 VarPTm) B0 ≡ subst_PTm (scons a1 VarPTm) B1 ∈ PUniv i.
+Lemma SSu_Pi_Proj2 n Γ (a0 a1 A0 A1 : PTm n) 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.
 Proof.
-  move /SemEq_SemWt => [hA][hB]/DJoin.bind_inj [_ [h1 h2]] /SemEq_SemWt [ha0][ha1]ha.
-  move : h2 => [B [h2 h2']].
-  move : ha => [c [ha ha']].
-  apply SemWt_SemEq; eauto using SBind_inst.
-  apply SBind_inst with (p := p) (A := A1); eauto.
-  apply : ST_Conv_E; eauto. hauto l:on use:SBind_inv1.
-  rewrite /DJoin.R.
-  eexists. split.
-  apply : relations.rtc_transitive.
-  apply REReds.substing. apply h2.
-  apply REReds.cong.
-  have : forall i0, rtc RERed.R (scons a0 VarPTm i0) (scons c VarPTm i0) by hauto q:on ctrs:rtc inv:option.
-  apply.
-  apply : relations.rtc_transitive. apply REReds.substing; eauto.
-  apply REReds.cong.
-  hauto lq:on ctrs:rtc inv:option.
+  move /SemLEq_SemWt => [/Sub.bind_inj [_ [h1 h2]]].
+  move => [i][hP0]hP1 /SemEq_SemWt [ha0][ha1]ha.
+  apply : SemWt_SemLEq; eauto using SBind_inst;
+    last by hauto l:on use:Sub.cong.
+  apply SBind_inst with (p := PPi) (A := A0); eauto.
+  apply : ST_Conv'; eauto. hauto l:on use:SBind_inv1.
+Qed.
+
+Lemma SSu_Sig_Proj2 n Γ (a0 a1 A0 A1 : PTm n) 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.
+Proof.
+  move /SemLEq_SemWt => [/Sub.bind_inj [_ [h1 h2]]].
+  move => [i][hP0]hP1 /SemEq_SemWt [ha0][ha1]ha.
+  apply : SemWt_SemLEq; eauto using SBind_inst;
+    last by hauto l:on use:Sub.cong.
+  apply SBind_inst with (p := PSig) (A := A1); eauto.
+  apply : ST_Conv'; eauto. hauto l:on use:SBind_inv1.
 Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SE_Bind_Proj1 SE_Bind_Proj2 SemWff_nil SemWff_cons : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2  SemWff_nil SemWff_cons : sem.

From 0c83044d724dd741db19f61bd0e62e87abd81337 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 22:45:07 -0500
Subject: [PATCH 061/210] Update the typing rules with projections

---
 theories/logrel.v |  1 +
 theories/typing.v | 28 ++++++++++++++++++++++------
 2 files changed, 23 insertions(+), 6 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 70c0706..aa47522 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -24,6 +24,7 @@ Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop)
 | InterpExt_Ne A :
   SNe A ->
   ⟦ A ⟧ i ;; I ↘ (fun a => exists v, rtc TRedSN a v /\ SNe v)
+
 | InterpExt_Bind p A B PA PF :
   ⟦ A ⟧ i ;; I ↘ PA ->
   (forall a, PA a -> exists PB, PF a PB) ->
diff --git a/theories/typing.v b/theories/typing.v
index 2d7078e..97d2e19 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -108,21 +108,37 @@ with LEq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   i <= j ->
   Γ ⊢ PUniv i : PTm n ≲ PUniv j
 
-| Su_Bind n Γ p (A0 A1 : PTm n) B0 B1 :
+| Su_Pi n Γ (A0 A1 : PTm n) B0 B1 :
+  ⊢ Γ ->
   Γ ⊢ A1 ≲ A0 ->
   funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≲ B1 ->
-  Γ ⊢ PBind p A0 B0 ≲ PBind p A1 B1
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1
+
+| Su_Sig n Γ (A0 A1 : PTm n) B0 B1 :
+  ⊢ Γ ->
+  Γ ⊢ A0 ≲ A1 ->
+  funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ≲ B1 ->
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1
 
 | Su_Eq n Γ (A : PTm n) B i  :
   Γ ⊢ A ≡ B ∈ PUniv i ->
   Γ ⊢ A ≲ B
 
-| Su_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 :
-  Γ ⊢ PBind p A0 B0 ≲ PBind p A1 B1 ->
+| Su_Pi_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
   Γ ⊢ A1 ≲ A0
 
-| Su_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 :
-  Γ ⊢ PBind p A0 B0 ≲ PBind p A1 B1 ->
+| Su_Sig_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
+  Γ ⊢ A0 ≲ A1
+
+| Su_Pi_Proj2 n Γ (a0 a1 A0 A1 : PTm n) 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 :
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
   Γ ⊢ a0 ≡ a1 ∈ A0 ->
   Γ ⊢ subst_PTm (scons a0 VarPTm) B0 ≲ subst_PTm (scons a1 VarPTm) B1
 

From 932662d5d90204e8adc3afb68f06a5d9115896a8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 22:52:50 -0500
Subject: [PATCH 062/210] Finish soundness proof

---
 theories/logrel.v    | 10 ++++++++--
 theories/soundness.v |  7 ++++---
 2 files changed, 12 insertions(+), 5 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index aa47522..1f48a97 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1120,7 +1120,7 @@ Proof.
   qauto l:on use:SemWt_SemLEq, Sub.transitive.
 Qed.
 
-Lemma ST_Univ n Γ i j :
+Lemma ST_Univ' n Γ i j :
   i < j ->
   Γ ⊨ PUniv i : PTm n ∈ PUniv j.
 Proof.
@@ -1128,6 +1128,12 @@ Proof.
   apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ.
 Qed.
 
+Lemma ST_Univ n Γ i :
+  Γ ⊨ PUniv i : PTm n ∈ PUniv (S i).
+Proof.
+  apply ST_Univ'. lia.
+Qed.
+
 Lemma SSu_Univ n Γ i j :
   i <= j ->
   Γ ⊨ PUniv i : PTm n ≲ PUniv j.
@@ -1227,4 +1233,4 @@ Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2  SemWff_nil SemWff_cons : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ : sem.
diff --git a/theories/soundness.v b/theories/soundness.v
index c2a40f0..8dd87da 100644
--- a/theories/soundness.v
+++ b/theories/soundness.v
@@ -5,7 +5,8 @@ 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).
-  apply wt_mutual; eauto with sem;[idtac].
-  hauto l:on use:SE_Pair.
+  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> Γ ⊨ a ≡ b ∈ A) /\
+  (forall n Γ (a b : PTm n), Γ ⊢ a ≲ b -> Γ ⊨ a ≲ b).
+  apply wt_mutual; eauto with sem; [hauto l:on use:SE_Pair].
+  Unshelve. all : exact 0.
 Qed.

From 439678670106a9f7c8abba8aeff0c53502a5a1d3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 8 Feb 2025 22:56:45 -0500
Subject: [PATCH 063/210] Reformulate renaming

---
 theories/structural.v | 14 +++++---------
 theories/typing.v     |  4 +++-
 2 files changed, 8 insertions(+), 10 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index 93fa17c..2423bc0 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -3,17 +3,11 @@ From Hammer Require Import Tactics.
 Require Import ssreflect.
 
 
-Lemma lem :
-  (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 -> ...).
-Proof. apply wt_mutual. ...
-
-
 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 A : PTm n), Γ ⊢ a ≡ b ∈ A -> ⊢ Γ) /\
+  (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> ⊢ Γ).
 Proof. apply wt_mutual; eauto. Qed.
 
 #[export]Hint Constructors Wt Wff Eq : wt.
@@ -35,7 +29,9 @@ Lemma renaming :
   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> forall m Δ (ξ : fin n -> fin m), ⊢ Δ -> 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 Δ Γ ξ ->
-     Δ ⊢ ren_PTm ξ a ≡ ren_PTm ξ b ∈ ren_PTm ξ A).
+     Δ ⊢ 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 Δ Γ ξ ->
+    Δ ⊢ ren_PTm ξ A ≲ ren_PTm ξ B).
 Proof.
   apply wt_mutual => //=; eauto 3 with wt.
   - move => n Γ i hΓ _ m Δ ξ hΔ hξ.
diff --git a/theories/typing.v b/theories/typing.v
index 97d2e19..4e5f60f 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -105,6 +105,7 @@ with LEq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   Γ ⊢ A ≲ C
 
 | Su_Univ n Γ i j :
+  ⊢ Γ ->
   i <= j ->
   Γ ⊢ PUniv i : PTm n ≲ PUniv j
 
@@ -164,5 +165,6 @@ Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind, le_ind.
 (* Lemma lem : *)
 (*   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> ...) /\ *)
 (*   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> ...)  /\ *)
-(*   (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> ...). *)
+(*   (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> ...) /\ *)
+(*   (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> ...). *)
 (* Proof. apply wt_mutual. ... *)

From ab1bd8eef8e8225e7ca6969a82f602514a9c2d91 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 9 Feb 2025 16:12:57 -0500
Subject: [PATCH 064/210] Add semantic rules for function beta and eta

---
 theories/logrel.v     | 56 ++++++++++++++++++++++++++----
 theories/structural.v | 79 +++++++++++++++++++++++++++++++++++++++++--
 theories/typing.v     | 13 +++----
 3 files changed, 132 insertions(+), 16 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 1f48a97..d68aebb 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -755,6 +755,13 @@ Proof.
   asimpl. hauto lq:on.
 Qed.
 
+Lemma ST_App' n Γ (b a : PTm n) 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 :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊨ a ∈ A ->
@@ -1010,6 +1017,48 @@ Proof.
   apply SemWt_SemEq; eauto using DJoin.AbsCong, ST_Abs.
 Qed.
 
+Lemma SBind_inv1 n Γ i p (A : PTm n) 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 :
+  ⊨ Γ ->
+  Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
+  Γ ⊨ b ∈ PBind PPi A B ->
+  Γ ⊨ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ∈ PBind PPi A B.
+Proof.
+  move => hΓ h0 h1. 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.
+  2 : {
+    apply ST_Var.
+    eauto using SemWff_cons.
+  }
+  change (PBind PPi (ren_PTm shift A) (ren_PTm (upRen_PTm_PTm shift) B)) with
+    (ren_PTm shift (PBind PPi A B)).
+  apply : weakening_Sem; eauto.
+Qed.
+
+Lemma SE_AppAbs n Γ (a : PTm (S n)) b A B i:
+  Γ ⊨ PBind PPi A B ∈ PUniv i ->
+  Γ ⊨ b ∈ A ->
+  funcomp (ren_PTm shift) (scons 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.
+  move => ρ hρ.
+  have {}/h0 := hρ.
+  move => [k][PA][hPA]hb.
+  move : ρ_ok_cons hPA hb (hρ); repeat move/[apply].
+  move => {}/h1.
+  by asimpl.
+  apply DJoin.FromRRed0.
+  apply RRed.AppAbs.
+Qed.
+
 Lemma SE_Conv' n Γ (a b : PTm n) A B i :
   Γ ⊨ a ≡ b ∈ A ->
   Γ ⊨ B ∈ PUniv i ->
@@ -1094,13 +1143,6 @@ Proof.
   apply : ST_Conv_E; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
 Qed.
 
-Lemma SBind_inv1 n Γ i p (A : PTm n) 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 SSu_Eq n Γ (A B : PTm n) i :
   Γ ⊨ A ≡ B ∈ PUniv i ->
   Γ ⊨ A ≲ B.
diff --git a/theories/structural.v b/theories/structural.v
index 2423bc0..3882d4c 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -24,6 +24,71 @@ Proof.
   by asimpl.
 Qed.
 
+Lemma Su_Wt n Γ a i :
+  Γ ⊢ a ∈ @PUniv n i ->
+  Γ ⊢ a ≲ a.
+Proof. hauto lq:on ctrs:LEq, Eq. Qed.
+
+Lemma Wt_Univ n Γ a A i
+  (h : Γ ⊢ a ∈ A) :
+  Γ ⊢ @PUniv n i ∈ PUniv (S i).
+Proof.
+  hauto lq:on ctrs:Wt use:wff_mutual.
+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 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 /\ Γ ⊢ A ∈ PUniv i). *)
+(* Proof. *)
+(*   apply wt_mutual => //=. *)
+(*   - admit. *)
+(*   - admit. *)
+
+Lemma T_App' n Γ (b a : PTm n) 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 :
+  U = subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ b ∈ U ->
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ PPair a b ∈ PBind PSig A B.
+Proof.
+  move => ->. eauto using T_Pair.
+Qed.
+
+Lemma T_Proj2' n Γ (a : PTm n) 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_Bind' n Γ i p (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons 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 renaming :
   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> True) /\
   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> forall m Δ (ξ : fin n -> fin m), ⊢ Δ -> renaming_ok Δ Γ ξ ->
@@ -40,6 +105,14 @@ Proof.
   - hauto lq:on rew:off ctrs:Wt, Wff  use:renaming_up.
   - move => n Γ a A B i hP ihP ha iha m Δ ξ hΔ hξ.
     apply : T_Abs; eauto.
-    apply iha; last by apply renaming_up.
-    econstructor; eauto.
-    by apply renaming_up.
+    move : ihP(hΔ) (hξ); repeat move/[apply]. move/Pi_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Δ.
+    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.
+  - hauto lq:on ctrs:Wt,LEq.
+  - hauto lq:on ctrs:Eq.
+  - move => n Γ i p A0 A1 B0 B1 hΓ _ hA ihA hB ihB m Δ ξ hΔ hξ.
+    apply E_Bind'; eauto.
+    apply ihB; last by hauto l:on use:renaming_up.
diff --git a/theories/typing.v b/theories/typing.v
index 4e5f60f..2156d61 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -9,10 +9,10 @@ Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   ⊢ Γ ->
   Γ ⊢ VarPTm i ∈ Γ i
 
-| T_Bind n Γ i j p (A : PTm n) (B : PTm (S n)) :
+| T_Bind n Γ i p (A : PTm n) (B : PTm (S n)) :
   Γ ⊢ A ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv j ->
-  Γ ⊢ PBind p A B ∈ PUniv (max i j)
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i ->
+  Γ ⊢ PBind p A B ∈ PUniv i
 
 | T_Abs n Γ (a : PTm (S n)) A B i :
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
@@ -61,11 +61,12 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ b ≡ c ∈ A ->
   Γ ⊢ a ≡ c ∈ A
 
-| E_Bind n Γ i j p (A0 A1 : PTm n) B0 B1 :
+| E_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
   ⊢ Γ ->
+  Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
-  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv j ->
-  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv (max i j)
+  funcomp (ren_PTm shift) (scons 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 :
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->

From 133968dd230a106bd1b2b1c1b5885db9f4e792f2 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 9 Feb 2025 16:25:34 -0500
Subject: [PATCH 065/210] Add semantic eta laws for pair

---
 theories/logrel.v | 35 +++++++++++++++++++++++++++++++++++
 1 file changed, 35 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index d68aebb..b944aea 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1131,6 +1131,41 @@ Proof.
   hauto lq:on use: DJoin.cong,  DJoin.ProjCong.
 Qed.
 
+Lemma SE_ProjPair1 n Γ (a b : PTm n) 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.
+Proof.
+  move => h0 h1 h2.
+  apply SemWt_SemEq; eauto using ST_Proj1, ST_Pair.
+  apply DJoin.FromRRed0. apply RRed.ProjPair.
+Qed.
+
+Lemma SE_ProjPair2 n Γ (a b : PTm n) 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.
+Proof.
+  move => h0 h1 h2.
+  apply SemWt_SemEq; eauto using ST_Proj2, ST_Pair.
+  apply : ST_Conv_E. apply : ST_Proj2; eauto. apply : ST_Pair; eauto.
+  hauto l:on use:SBind_inst.
+  apply DJoin.cong. apply DJoin.FromRRed0. apply RRed.ProjPair.
+  apply DJoin.FromRRed0. apply RRed.ProjPair.
+Qed.
+
+Lemma SE_PairEta n Γ (a : PTm n) 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.
+Proof.
+  move => h0 h. apply SemWt_SemEq; eauto.
+  apply : ST_Pair; eauto using ST_Proj1, ST_Proj2.
+  rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
+Qed.
+
 Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) A B :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊨ b0 ≡ b1 ∈ PBind PPi A B ->

From 5b925e3fa1d10407523b832c5b44b163892e78aa Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 9 Feb 2025 16:33:43 -0500
Subject: [PATCH 066/210] Add beta and eta rules to syntactic typing

---
 theories/logrel.v |  8 +++++---
 theories/typing.v | 37 +++++++++++++++++++++++++++++++++++++
 2 files changed, 42 insertions(+), 3 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index b944aea..8248063 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1028,9 +1028,10 @@ Lemma SE_AppEta n Γ (b : PTm n) A B i :
   ⊨ Γ ->
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊨ b ∈ PBind PPi A B ->
-  Γ ⊨ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ∈ PBind PPi A B.
+  Γ ⊨ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B.
 Proof.
-  move => hΓ h0 h1. apply : ST_Abs; eauto.
+  move => hΓ 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.
   2 : {
@@ -1040,6 +1041,7 @@ Proof.
   change (PBind PPi (ren_PTm shift A) (ren_PTm (upRen_PTm_PTm shift) B)) with
     (ren_PTm shift (PBind PPi A B)).
   apply : weakening_Sem; eauto.
+  hauto q:on ctrs:rtc,RERed.R.
 Qed.
 
 Lemma SE_AppAbs n Γ (a : PTm (S n)) b A B i:
@@ -1310,4 +1312,4 @@ Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta : sem.
diff --git a/theories/typing.v b/theories/typing.v
index 2156d61..2733b9c 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -48,6 +48,7 @@ Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   Γ ⊢ a ∈ B
 
 with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
+(* Structural *)
 | E_Refl n Γ (a : PTm n) A :
   Γ ⊢ a ∈ A ->
   Γ ⊢ a ≡ a ∈ A
@@ -61,6 +62,7 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ b ≡ c ∈ A ->
   Γ ⊢ a ≡ c ∈ A
 
+(* Congruence *)
 | E_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
   ⊢ Γ ->
   Γ ⊢ A0 ∈ PUniv i ->
@@ -99,12 +101,45 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ A ≲ B ->
   Γ ⊢ a ≡ b ∈ B
 
+(* Beta *)
+| E_AppAbs n Γ (a : PTm (S n)) b A B i:
+  Γ ⊢ PBind PPi A B ∈ PUniv i ->
+  Γ ⊢ b ∈ A ->
+  funcomp (ren_PTm shift) (scons 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 :
+  Γ ⊢ 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 :
+  Γ ⊢ 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
+
+(* Eta *)
+| E_AppEta n Γ (b : PTm n) 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 :
+  Γ ⊢ 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 :=
+(* Structural *)
 | Su_Transitive  n Γ (A B C : PTm n) :
   Γ ⊢ A ≲ B ->
   Γ ⊢ B ≲ C ->
   Γ ⊢ A ≲ C
 
+(* Congruence *)
 | Su_Univ n Γ i j :
   ⊢ Γ ->
   i <= j ->
@@ -122,10 +157,12 @@ with LEq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ≲ B1 ->
   Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1
 
+(* Injecting from equalities *)
 | Su_Eq n Γ (A : PTm n) B i  :
   Γ ⊢ A ≡ B ∈ PUniv i ->
   Γ ⊢ A ≲ B
 
+(* Projection axioms *)
 | Su_Pi_Proj1 n Γ (A0 A1 : PTm n) B0 B1 :
   Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
   Γ ⊢ A1 ≲ A0

From 20bf38a3cab22fa14dee83a020e3af3bdbfb7b1e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 9 Feb 2025 16:46:17 -0500
Subject: [PATCH 067/210] Fix the fundamental theorem yet again

---
 theories/logrel.v | 36 ++++++++++++++++++++++++++++--------
 1 file changed, 28 insertions(+), 8 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 8248063..52a4438 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -702,7 +702,7 @@ Proof.
 Qed.
 
 
-Lemma ST_Bind n Γ i j p (A : PTm n) (B : PTm (S n)) :
+Lemma ST_Bind' n Γ i j p (A : PTm n) (B : PTm (S n)) :
   Γ ⊨ A ∈ PUniv i ->
   funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv j ->
   Γ ⊨ PBind p A B ∈ PUniv (max i j).
@@ -717,6 +717,17 @@ 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)) :
+  Γ ⊨ A ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊨ B ∈ PUniv i ->
+  Γ ⊨ PBind p A B ∈ PUniv i.
+Proof.
+  move => h0 h1.
+  replace i with (max i i) by lia.
+  move : h0 h1.
+  apply ST_Bind'.
+Qed.
+
 Lemma ST_Abs n Γ (a : PTm (S n)) A B i :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
   funcomp (ren_PTm shift) (scons A Γ) ⊨ a ∈ B ->
@@ -990,7 +1001,7 @@ Lemma Γ_eq_cons' n (Γ : fin n -> PTm n) 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' n Γ i j p (A0 A1 : PTm n) B0 B1 :
   ⊨ Γ ->
   Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
   funcomp (ren_PTm shift) (scons A0 Γ) ⊨ B0 ≡ B1 ∈ PUniv j ->
@@ -999,8 +1010,8 @@ Proof.
   move => hΓ 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.
+  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.
@@ -1008,6 +1019,15 @@ Proof.
   hauto lq:on use:Γ_eq_cons'.
 Qed.
 
+Lemma SE_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
+  ⊨ Γ ->
+  Γ ⊨ A0 ≡ A1 ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons 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 :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
   funcomp (ren_PTm shift) (scons A Γ) ⊨ a ≡ b ∈ B ->
@@ -1234,8 +1254,8 @@ Proof.
   move : hA => [hA0][i][hA1]hA2.
   move : hB => [hB0][j][hB1]hB2.
   apply : SemWt_SemLEq; last by hauto l:on use:Sub.PiCong.
-  hauto l:on use:ST_Bind.
-  apply ST_Bind; eauto.
+  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.
   move => ρ hρ.
   suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
@@ -1256,8 +1276,8 @@ Proof.
   move : hA => [hA0][i][hA1]hA2.
   move : hB => [hB0][j][hB1]hB2.
   apply : SemWt_SemLEq; last by hauto l:on use:Sub.SigCong.
-  2 : { hauto l:on use:ST_Bind. }
-  apply ST_Bind; eauto.
+  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.
   move => ρ hρ.

From df5c6bf0b161b014fe4b038450a74e601982ead0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 9 Feb 2025 17:05:36 -0500
Subject: [PATCH 068/210] Minor

---
 theories/structural.v | 6 +++++-
 1 file changed, 5 insertions(+), 1 deletion(-)

diff --git a/theories/structural.v b/theories/structural.v
index 3882d4c..17dfcb1 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -84,6 +84,7 @@ Lemma T_Proj2' n Γ (a : PTm n) A B U :
 Proof. move => ->. apply T_Proj2. Qed.
 
 Lemma E_Bind' n Γ i p (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
   funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i.
@@ -113,6 +114,9 @@ Proof.
   - move => n Γ a A B ha iha m Δ ξ hΔ hξ. apply : T_Proj2'; eauto. by asimpl.
   - hauto lq:on ctrs:Wt,LEq.
   - hauto lq:on ctrs:Eq.
-  - move => n Γ i p A0 A1 B0 B1 hΓ _ hA ihA hB ihB m Δ ξ hΔ hξ.
+  - move => n Γ i p A0 A1 B0 B1 hΓ _ hA ihA hA' ihA' hB ihB m Δ ξ hΔ hξ.
     apply E_Bind'; eauto.
     apply ihB; last by hauto l:on use:renaming_up.
+    hauto lq:on ctrs:Wff,Wt use:renaming_up.
+  - move => n Γ a b A B i hP ihP ha iha m Δ ξ hΔ hξ.
+    apply : E_Abs; eauto. apply iha; last by hauto l:on use:renaming_up.

From 881b2e38250f73ef06d0a1a8c76faca54ce99af6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Sun, 9 Feb 2025 17:05:43 -0500
Subject: [PATCH 069/210] Add missing premise

---
 theories/typing.v | 1 +
 1 file changed, 1 insertion(+)

diff --git a/theories/typing.v b/theories/typing.v
index 2733b9c..7269496 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -71,6 +71,7 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i
 
 | E_Abs n Γ (a b : PTm (S n)) A B i :
+  Γ ⊢ A ∈ PUniv i ->
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
   funcomp (ren_PTm shift) (scons A Γ) ⊢ a ≡ b ∈ B ->
   Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B

From ea1fba8ae9d299dfdcd1abfbaee6b74a0bf55b8d Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 9 Feb 2025 20:41:27 -0500
Subject: [PATCH 070/210] Finish syntactic renaming

---
 theories/structural.v | 119 +++++++++++++++++++++++++++++++++++++++---
 theories/typing.v     |   7 +--
 2 files changed, 117 insertions(+), 9 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index 17dfcb1..f15c485 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -50,6 +50,19 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
+Lemma Sig_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 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)  /\ *)
@@ -83,6 +96,15 @@ Lemma T_Proj2' n Γ (a : PTm n) A B U :
   Γ ⊢ PProj PR a ∈ U.
 Proof. move => ->. apply T_Proj2. Qed.
 
+Lemma E_Proj2' n Γ i (a b : PTm n) 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 :
   Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
@@ -90,6 +112,54 @@ Lemma E_Bind' n Γ i p (A0 A1 : PTm n) B0 B1 :
   Γ ⊢ 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 :
+  U = subst_PTm (scons a0 VarPTm) B ->
+  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+  Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B ->
+  Γ ⊢ a0 ≡ a1 ∈ A ->
+  Γ ⊢ 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 :
+  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 ->
+  Γ ⊢ PApp (PAbs a) b ≡ u ∈ U.
+  move => -> ->. apply E_AppAbs. Qed.
+
+Lemma E_ProjPair2' n Γ (a b : PTm n) A B i U :
+  U = subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+  Γ ⊢ a ∈ A ->
+  Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B ->
+  Γ ⊢ PProj PR (PPair a b) ≡ b ∈ U.
+Proof. move => ->. apply E_ProjPair2. Qed.
+
+Lemma E_AppEta' n Γ (b : PTm n) 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 :
+  U = subst_PTm (scons a0 VarPTm) B0 ->
+  T = subst_PTm (scons a1 VarPTm) B1 ->
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
+  Γ ⊢ a0 ≡ a1 ∈ A1 ->
+  Γ ⊢ U ≲ T.
+Proof. move => -> ->. apply Su_Pi_Proj2. Qed.
+
+Lemma Su_Sig_Proj2' n Γ (a0 a1 A0 A1 : PTm n) 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 ->
+  Γ ⊢ a0 ≡ a1 ∈ A0 ->
+  Γ ⊢ U ≲ T.
+Proof. move => -> ->. apply Su_Sig_Proj2. 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 Δ Γ ξ ->
@@ -114,9 +184,46 @@ Proof.
   - move => n Γ a A B ha iha m Δ ξ hΔ hξ. apply : T_Proj2'; eauto. by asimpl.
   - hauto lq:on ctrs:Wt,LEq.
   - hauto lq:on ctrs:Eq.
-  - move => n Γ i p A0 A1 B0 B1 hΓ _ hA ihA hA' ihA' hB ihB m Δ ξ hΔ hξ.
-    apply E_Bind'; eauto.
-    apply ihB; last by hauto l:on use:renaming_up.
-    hauto lq:on ctrs:Wff,Wt use:renaming_up.
-  - move => n Γ a b A B i hP ihP ha iha m Δ ξ hΔ hξ.
-    apply : E_Abs; eauto. apply iha; last by hauto l:on use:renaming_up.
+  - 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 : ihPi (hΔ) (hξ). repeat move/[apply].
+    move => /Pi_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ξ.
+    apply : E_Pair; eauto.
+    move : ihb hΔ hξ. repeat move/[apply].
+    by asimpl.
+  - move => *. apply : E_Proj2'; eauto. 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 : ihP (hξ) (hΔ). repeat move/[apply].
+    move /Pi_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 : {hP} ihP (hξ) (hΔ). repeat move/[apply].
+    move /Sig_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ξ.
+    apply : E_ProjPair2'; eauto. by asimpl.
+    move : ihb hξ hΔ; repeat move/[apply]. by asimpl.
+  - move => *.
+    apply : E_AppEta'; eauto. by asimpl.
+  - qauto l:on use:E_PairEta.
+  - hauto lq:on ctrs:LEq.
+  - qauto l:on ctrs:LEq.
+  - hauto lq:on ctrs:Wff use:renaming_up, Su_Pi.
+  - hauto lq:on ctrs:Wff use:Su_Sig, renaming_up.
+  - hauto q:on ctrs:LEq.
+  - hauto lq:on ctrs:LEq.
+  - qauto l:on ctrs:LEq.
+  - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
+  - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
+Qed.
diff --git a/theories/typing.v b/theories/typing.v
index 7269496..052eab8 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -71,7 +71,6 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i
 
 | E_Abs n Γ (a b : PTm (S n)) A B i :
-  Γ ⊢ A ∈ PUniv i ->
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
   funcomp (ren_PTm shift) (scons A Γ) ⊢ a ≡ b ∈ B ->
   Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B
@@ -146,14 +145,16 @@ with LEq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   i <= j ->
   Γ ⊢ PUniv i : PTm n ≲ PUniv j
 
-| Su_Pi n Γ (A0 A1 : PTm n) B0 B1 :
+| Su_Pi n Γ (A0 A1 : PTm n) B0 B1 i :
   ⊢ Γ ->
+  Γ ⊢ A0 ∈ PUniv i ->
   Γ ⊢ A1 ≲ A0 ->
   funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≲ B1 ->
   Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1
 
-| Su_Sig n Γ (A0 A1 : PTm n) B0 B1 :
+| Su_Sig n Γ (A0 A1 : PTm n) B0 B1 i :
   ⊢ Γ ->
+  Γ ⊢ A1 ∈ PUniv i ->
   Γ ⊢ A0 ≲ A1 ->
   funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ≲ B1 ->
   Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1

From 2c47d78ad616676294a4f22126cb9a5b45a1da39 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 9 Feb 2025 23:32:07 -0500
Subject: [PATCH 071/210] Add stub for morphing

---
 theories/structural.v | 103 ++++++++++++++++++++++++++++++++++++++----
 1 file changed, 93 insertions(+), 10 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index f15c485..d3ce673 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -63,16 +63,6 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-(* 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 /\ Γ ⊢ A ∈ PUniv i). *)
-(* Proof. *)
-(*   apply wt_mutual => //=. *)
-(*   - admit. *)
-(*   - admit. *)
-
 Lemma T_App' n Γ (b a : PTm n) A B U :
   U = subst_PTm (scons a VarPTm) B ->
   Γ ⊢ b ∈ PBind PPi A B ->
@@ -227,3 +217,96 @@ Proof.
   - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
 Qed.
+
+Definition morphing_ok {n m} Δ Γ (ρ : fin n -> PTm m) :=
+  forall i, Δ ⊢ ρ i ∈ subst_PTm ρ (Γ i).
+
+Lemma morphing_ren n m p Ξ Δ Γ
+  (ρ : fin n -> PTm m) (ξ : fin m -> fin p) :
+  ⊢ Ξ ->
+  renaming_ok Ξ Δ ξ -> morphing_ok Δ Γ ρ ->
+  morphing_ok Ξ Γ (funcomp (ren_PTm ξ) ρ).
+Proof.
+  move => hΞ hξ hρ.
+  move => i.
+  rewrite {1}/funcomp.
+  have -> :
+    subst_PTm (funcomp (ren_PTm ξ) ρ) (Γ i) =
+    ren_PTm ξ (subst_PTm ρ (Γ i)) by asimpl.
+  eapply renaming; eauto.
+Qed.
+
+Lemma morphing_ext n m Δ Γ (ρ : fin n -> PTm m) (a : PTm m) (A : PTm n)  :
+  morphing_ok Δ Γ ρ ->
+  Δ ⊢ a ∈ subst_PTm ρ A ->
+  morphing_ok
+  Δ (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ).
+Proof.
+  move => h ha i. destruct i as [i|]; by asimpl.
+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.
+Proof. sfirstorder use:renaming. Qed.
+
+Lemma renaming_wt' : forall n m Δ Γ a A (ξ : fin n -> fin m) u U,
+    u = ren_PTm ξ a -> U = ren_PTm ξ A ->
+    Γ ⊢ a ∈ A -> ⊢ Δ ->
+    renaming_ok Δ Γ ξ -> Δ ⊢ u ∈ U.
+Proof. hauto use:renaming_wt. 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 morphing_up n m Γ Δ (ρ : fin n -> PTm m) (A : PTm n) 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 ρ).
+Proof.
+  move => h h1 [:hp]. apply morphing_ext.
+  rewrite /morphing_ok.
+  move => i.
+  rewrite {2}/funcomp.
+  apply : renaming_wt'; 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.
+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 Δ Γ ρ ->
+     Δ ⊢ subst_PTm ρ a ∈ subst_PTm ρ A) /\
+  (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> forall m Δ (ρ : fin n -> PTm m), ⊢ Δ -> 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 Δ Γ ρ ->
+    Δ ⊢ subst_PTm ρ A ≲ subst_PTm ρ B).
+Proof.
+  apply wt_mutual => //=.
+  - best.
+
+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 /\ Γ ⊢ A ∈ 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].
+  -

From c8e84ef93c12290f6a38bfcbbfb9bbd965c03ae4 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 10 Feb 2025 00:17:01 -0500
Subject: [PATCH 072/210] Finish morphing

---
 theories/structural.v | 90 ++++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 88 insertions(+), 2 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index d3ce673..c9127d0 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -2,7 +2,6 @@ Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typin
 From Hammer Require Import Tactics.
 Require Import ssreflect.
 
-
 Lemma wff_mutual :
   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> True) /\
   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> ⊢ Γ)  /\
@@ -279,6 +278,31 @@ Proof.
   apply : T_Var';eauto. rewrite /funcomp. by asimpl.
 Qed.
 
+Lemma Wff_Cons' n Γ (A : PTm n) i :
+  Γ ⊢ A ∈ PUniv i ->
+  (* -------------------------------- *)
+  ⊢ funcomp (ren_PTm shift) (scons A Γ).
+Proof. hauto lq:on rew:off use:Wff_Cons, wff_mutual. Qed.
+
+Lemma weakening_wt : forall n Γ (a A B : PTm n) i,
+    Γ ⊢ B ∈ PUniv i ->
+    Γ ⊢ a ∈ A ->
+    funcomp (ren_PTm shift) (scons B Γ) ⊢ ren_PTm shift a ∈ ren_PTm shift A.
+Proof.
+  move => n Γ 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,
+    u = ren_PTm shift a ->
+    U = ren_PTm shift A ->
+    Γ ⊢ B ∈ PUniv i ->
+    Γ ⊢ a ∈ A ->
+    funcomp (ren_PTm shift) (scons 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 Δ Γ ρ ->
@@ -289,7 +313,69 @@ Lemma morphing :
     Δ ⊢ subst_PTm ρ A ≲ subst_PTm ρ B).
 Proof.
   apply wt_mutual => //=.
-  - best.
+  - hauto lq:on use:morphing_up, Wff_Cons', T_Bind.
+  - move => n Γ a A B i hP ihP ha iha m Δ ρ hΔ hρ.
+    move : ihP (hΔ) (hρ); repeat move/[apply].
+    move /Pi_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.
+    apply : T_Abs; eauto.
+    apply : iha.
+    hauto lq:on use:Wff_Cons', Pi_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Δ.
+    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.
+  - hauto lq:on ctrs:Wt,LEq.
+  - qauto l:on ctrs:Wt.
+  - hauto lq:on ctrs:Eq.
+  - 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 : ihPi (hΔ) (hρ). repeat move/[apply].
+    move => /Pi_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ρ.
+    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.
+  - qauto l:on ctrs:Eq, LEq.
+  - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ρ hΔ hρ.
+    move : ihP (hρ) (hΔ). repeat move/[apply].
+    move /Pi_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 : {hP} ihP (hρ) (hΔ). repeat move/[apply].
+    move /Sig_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ρ.
+    apply : E_ProjPair2'; eauto. by asimpl.
+    move : ihb hρ hΔ; repeat move/[apply]. by asimpl.
+  - move => *.
+    apply : E_AppEta'; eauto. by asimpl.
+  - qauto l:on use:E_PairEta.
+  - hauto lq:on ctrs:LEq.
+  - qauto l:on ctrs:LEq.
+  - hauto lq:on ctrs:Wff use:morphing_up, Su_Pi.
+  - hauto lq:on ctrs:Wff use:Su_Sig, morphing_up.
+  - hauto q:on ctrs:LEq.
+  - hauto lq:on ctrs:LEq.
+  - qauto l:on ctrs:LEq.
+  - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
+  - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
+Qed.
 
 Lemma regularity :
   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> forall i, exists j, Γ ⊢ Γ i ∈ PUniv j) /\

From 26e3c1c42a7246ec8a34f07ce357ed5ee491aaa6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 10 Feb 2025 01:15:44 -0500
Subject: [PATCH 073/210] Add some cases for regularity

---
 theories/structural.v | 85 ++++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 83 insertions(+), 2 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index c9127d0..e02be33 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -1,6 +1,7 @@
 Require Import Autosubst2.core Autosubst2.fintype 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) /\
@@ -377,11 +378,38 @@ 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.
+Proof. sfirstorder use:morphing. Qed.
+
+Lemma morphing_wt' : forall n m Δ Γ a A (ρ : fin n -> PTm m) 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.
+Proof.
+  move => n Γ hΓ.
+  rewrite /morphing_ok.
+  move => i. asimpl. by apply 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 ->
+    Γ ⊢ 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.
+  abstract : hΓ. sfirstorder use:wff_mutual.
+  apply morphing_ext; last by asimpl.
+  by apply morphing_id.
+Qed.
+
 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 /\ Γ ⊢ A ∈ PUniv i).
+  (forall n Γ (A B : PTm n), Γ ⊢ 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.
@@ -395,4 +423,57 @@ Proof.
     econstructor; eauto.
     apply : renaming_shift; eauto.
   - move => n Γ b a A B hb [i ihb] ha [j iha].
-  -
+    move /Pi_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:Sig_Inv.
+  - move => n Γ a A B ha [i /Sig_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.
+  - sfirstorder.
+  - sfirstorder.
+  - sfirstorder.
+  - move => n Γ 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.
+    admit.
+    eauto using T_Univ.
+  - hauto lq:on ctrs:Wt,Eq.
+  - 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.
+    apply : T_Conv; eauto.
+    qauto use:T_App.
+    apply Su_Pi_Proj2 with (A0 := A) (A1 := A).
+    apply : Su_Eq; apply E_Refl; eauto.
+    by apply E_Symmetric.
+    sauto lq:on use:Pi_Inv, substing_wt.
+  - admit.
+  - admit.
+  - admit.
+  - hauto lq:on ctrs:Wt.
+  - admit.
+  - admit.
+  - admit.
+  - hauto lq:on ctrs:Wt.
+  - move => n Γ 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)).
+    have [? ?] : S i <= S (Nat.max i j) /\ S j <= S (Nat.max i j) by lia.
+    hauto lq:on ctrs:Wt,LEq.
+  - admit.
+  - admit.
+  - sfirstorder.
+  - admit.
+  - admit.
+  - admit.
+  - admit.
+Admitted.

From 5918fdf47ea7d31a3473f521f0250b72527d3d43 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Feb 2025 13:51:35 -0500
Subject: [PATCH 074/210] Prove all but 5 cases of regularity

---
 theories/structural.v | 76 ++++++++++++++++++++++++++++++++++++-------
 1 file changed, 64 insertions(+), 12 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index e02be33..e1d8e2c 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -405,6 +405,45 @@ Proof.
   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 ->
+  Γ ⊢ A0 ∈ PUniv i ->
+  Γ ⊢ A1 ∈ PUniv j ->
+  Γ ⊢ A1 ≲ A0 ->
+  funcomp (ren_PTm shift) (scons A1 Γ) ⊢ a ∈ A.
+Proof.
+  move => h0 h1 h2 h3.
+  replace a with (subst_PTm VarPTm a); last by asimpl.
+  replace A with (subst_PTm VarPTm A); last by asimpl.
+  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.
+    by substify.
+  - move => [:hΓ'].
+    apply : T_Conv.
+    apply T_Var.
+    abstract : hΓ'.
+    eauto using Wff_Cons'.
+    rewrite /funcomp. asimpl. substify. asimpl.
+    renamify.
+    eapply renaming; eauto.
+    apply : renaming_shift; eauto.
+Qed.
+
+Lemma bind_inst n Γ p (A : PTm n) 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.
+Proof.
+  move => h h0.
+  have {}h : Γ ⊢ PBind p A B ≲ PBind p A B by eauto using E_Refl, Su_Eq.
+  case : p h => //=; hauto l:on use:Su_Pi_Proj2, Su_Sig_Proj2.
+Qed.
+
 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)  /\
@@ -440,7 +479,8 @@ Proof.
     move => hB [ihB0 [ihB1 [i2 ihB2]]].
     repeat split => //=.
     qauto use:T_Bind.
-    admit.
+    apply T_Bind; eauto.
+    apply : ctx_eq_subst_one; eauto using Su_Eq, E_Symmetric.
     eauto using T_Univ.
   - hauto lq:on ctrs:Wt,Eq.
   - move => n Γ i b0 b1 a0 a1 A B hP _ hb [ihb0 [ihb1 [i0 ihb2]]]
@@ -449,17 +489,29 @@ Proof.
     qauto use:T_App.
     apply : T_Conv; eauto.
     qauto use:T_App.
-    apply Su_Pi_Proj2 with (A0 := A) (A1 := A).
-    apply : Su_Eq; apply E_Refl; eauto.
-    by apply E_Symmetric.
-    sauto lq:on use:Pi_Inv, substing_wt.
-  - admit.
-  - admit.
-  - admit.
+    move /E_Symmetric in ha.
+    by eauto using bind_inst.
+    hauto lq:on ctrs:Wt,Eq,LEq lq:on use:Pi_Inv, substing_wt.
+  - hauto lq:on use:bind_inst db:wt.
+  - hauto lq:on use:Sig_Inv db:wt.
+  - move => n Γ 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:Sig_Inv, substing_wt.
   - hauto lq:on ctrs:Wt.
-  - admit.
-  - admit.
-  - admit.
+  - hauto q:on use:substing_wt db:wt.
+  - hauto l:on use:bind_inst db:wt.
+  - move => n Γ b A B i hΓ ihΓ 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)
+                        by hauto lq:on use:weakening_wt, Pi_Inv.
+    apply : T_Abs; eauto.
+    apply : T_App'; eauto; rewrite-/ren_PTm.
+    by asimpl.
+    apply T_Var. sfirstorder use:wff_mutual.
   - hauto lq:on ctrs:Wt.
   - move => n Γ A B C hA [i [ihA0 ihA1]] hC [j [ihC0 ihC1]].
     have ? : ⊢ Γ by sfirstorder use:wff_mutual.
@@ -469,7 +521,7 @@ Proof.
   - move => n Γ 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.
-  - admit.
+  - best use:bind_inst.
   - admit.
   - sfirstorder.
   - admit.

From 8f8f428562238a33a9f7f61185ed32c5ce62ca56 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Feb 2025 14:16:43 -0500
Subject: [PATCH 075/210] Finish preservation

---
 theories/structural.v | 77 +++++++++++++++++++++++++++++++++++++++----
 1 file changed, 70 insertions(+), 7 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index e1d8e2c..8a6348e 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -444,6 +444,28 @@ Proof.
   case : p h => //=; hauto l:on use:Su_Pi_Proj2, Su_Sig_Proj2.
 Qed.
 
+Lemma Cumulativity n Γ (a : PTm n) i j :
+  i <= j ->
+  Γ ⊢ a ∈ PUniv i ->
+  Γ ⊢ a ∈ PUniv j.
+Proof.
+  move => h0 h1. apply : T_Conv; eauto.
+  apply Su_Univ => //=.
+  sfirstorder use:wff_mutual.
+Qed.
+
+Lemma T_Bind' n Γ i j p (A : PTm n) (B : PTm (S n)) :
+  Γ ⊢ A ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv j ->
+  Γ ⊢ PBind p A B ∈ PUniv (max i j).
+Proof.
+  move => h0 h1.
+  have [*] : i <= max i j /\ j <= max i j by lia.
+  qauto l:on ctrs:Wt use:Cumulativity.
+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)  /\
@@ -521,11 +543,52 @@ Proof.
   - move => n Γ 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.
-  - best use:bind_inst.
-  - admit.
+  - move => n Γ A0 A1 B0 B1 i hΓ ihΓ 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.
+    exists (max i0 i1). repeat split; eauto with wt.
+    apply T_Bind'; eauto.
+    sfirstorder use:ctx_eq_subst_one.
   - sfirstorder.
-  - admit.
-  - admit.
-  - admit.
-  - admit.
-Admitted.
+  - move => n Γ A0 A1 B0 B1 _ [i][ih0 ih1].
+    move /Pi_Inv : ih0 => [i0][h _].
+    move /Pi_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 /Sig_Inv : ih0 => [i0][h _].
+    move /Sig_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 => [i][ihP0]ihP1.
+    move => ha [iha0][iha1][j]ihA1.
+    move /Pi_Inv :ihP0 => [i0][ih0][ih0' _].
+    move /Pi_Inv :ihP1 => [i1][ih1][ih1' _].
+    have [*] : i0 <= max i0 i1 /\ i1 <= max i0 i1 by lia.
+    exists (max i0 i1).
+    split.
+    + apply Cumulativity with (i := i0); eauto.
+      have : Γ ⊢ a0 ∈ A0 by eauto using T_Conv.
+      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 => [i][ihP0]ihP1.
+    move => ha [iha0][iha1][j]ihA1.
+    move /Sig_Inv :ihP0 => [i0][ih0][ih0' _].
+    move /Sig_Inv :ihP1 => [i1][ih1][ih1' _].
+    have [*] : i0 <= max i0 i1 /\ i1 <= max i0 i1 by lia.
+    exists (max i0 i1).
+    split.
+    + apply Cumulativity with (i := i0); eauto.
+      move : substing_wt iha0 ih0';repeat move/[apply]. by asimpl.
+    + apply Cumulativity with (i := i1); eauto.
+      have : Γ ⊢ a1 ∈ A1 by eauto using T_Conv.
+      move : substing_wt ih1';repeat move/[apply]. by asimpl.
+Qed.

From 8105b5c410e347035f393d1b660e720b40fe6a1f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Feb 2025 17:01:40 -0500
Subject: [PATCH 076/210] Add admissible simple rules

---
 theories/admissible.v | 50 +++++++++++++++++++++++++++++++++++++++++++
 theories/common.v     | 46 ++++++++++++++++++++++++++++++++++++++-
 theories/fp_red.v     | 31 +--------------------------
 3 files changed, 96 insertions(+), 31 deletions(-)
 create mode 100644 theories/admissible.v

diff --git a/theories/admissible.v b/theories/admissible.v
new file mode 100644
index 0000000..347529e
--- /dev/null
+++ b/theories/admissible.v
@@ -0,0 +1,50 @@
+Require Import Autosubst2.core Autosubst2.fintype 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.
+
+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 ->
+  Γ ⊢ 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.
+Qed.
+
+Lemma E_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
+  Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i.
+Proof.
+  move => h0 h1.
+  have : Γ ⊢ A0 ∈ PUniv i by hauto l:on use:regularity.
+  have : ⊢ Γ by sfirstorder use:wff_mutual.
+  move : E_Bind h0 h1; repeat move/[apply].
+  firstorder.
+Qed.
diff --git a/theories/common.v b/theories/common.v
index d345d49..02c0b83 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -1,3 +1,47 @@
-Require Import Autosubst2.fintype Autosubst2.syntax.
+Require Import Autosubst2.fintype Autosubst2.syntax 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).
+
+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 q:on depth:2))
+  | _ => solve_anti_ren ()
+  end.
+
+Local Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
+
+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.
+Proof.
+  sblast inv:option.
+Qed.
+
+Lemma ren_injective n m (a b : PTm n) (ξ : fin n -> fin m) :
+  (forall i j, ξ i = ξ j -> i = j) ->
+  ren_PTm ξ a = ren_PTm ξ b ->
+  a = b.
+Proof.
+  move : m ξ b.
+  elim : n / a => //; try solve_anti_ren.
+
+  move => n a iha m ξ []//=.
+  move => u hξ [h].
+  apply iha in h. by subst.
+  destruct i, j=>//=.
+  hauto l:on.
+
+  move => n p A ihA B ihB m ξ []//=.
+  move => b A0 B0  hξ [?]. subst.
+  move => ?. have ? : A0 = A by firstorder. subst.
+  move => ?. have : B = B0. apply : ihB; eauto.
+  sauto.
+  congruence.
+Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 9998924..b906b6a 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.
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common.
 Require Import Btauto.
 Require Import Cdcl.Itauto.
 
@@ -1408,35 +1408,6 @@ Module ERed.
   (*   destruct a. *)
   (*   exact (ξ f). *)
 
-  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.
-  Proof.
-    sblast inv:option.
-  Qed.
-
-  Lemma ren_injective n m (a b : PTm n) (ξ : fin n -> fin m) :
-    (forall i j, ξ i = ξ j -> i = j) ->
-    ren_PTm ξ a = ren_PTm ξ b ->
-    a = b.
-  Proof.
-    move : m ξ b.
-    elim : n / a => //; try solve_anti_ren.
-
-    move => n a iha m ξ []//=.
-    move => u hξ [h].
-    apply iha in h. by subst.
-    destruct i, j=>//=.
-    hauto l:on.
-
-    move => n p A ihA B ihB m ξ []//=.
-    move => b A0 B0  hξ [?]. subst.
-    move => ?. have ? : A0 = A by firstorder. subst.
-    move => ?. have : B = B0. apply : ihB; eauto.
-    sauto.
-    congruence.
-  Qed.
-
   Lemma AppEta' n a u :
     u = (@PApp (S n) (ren_PTm shift a) (VarPTm var_zero)) ->
     R (PAbs u) a.

From bccf6eb860aa466b4d69ae9c0613405216dc509f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Feb 2025 18:40:42 -0500
Subject: [PATCH 077/210] Add Coquand's algorithm

---
 theories/admissible.v  |  19 ++++++++
 theories/algorithmic.v | 105 +++++++++++++++++++++++++++++++++++++++++
 theories/common.v      |  45 ++++++++++++++++++
 theories/fp_red.v      |  35 --------------
 4 files changed, 169 insertions(+), 35 deletions(-)
 create mode 100644 theories/algorithmic.v

diff --git a/theories/admissible.v b/theories/admissible.v
index 347529e..392e3cf 100644
--- a/theories/admissible.v
+++ b/theories/admissible.v
@@ -48,3 +48,22 @@ Proof.
   move : E_Bind h0 h1; repeat move/[apply].
   firstorder.
 Qed.
+
+Lemma E_App n Γ (b0 b1 a0 a1 : PTm n) A B :
+  Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B ->
+  Γ ⊢ a0 ≡ a1 ∈ A ->
+  Γ ⊢ PApp b0 a0 ≡ PApp b1 a1 ∈ subst_PTm (scons a0 VarPTm) B.
+Proof.
+  move => h.
+  have [i] : exists i, Γ ⊢ PBind PPi A B ∈ PUniv i by hauto l:on use:regularity.
+  move : E_App h. by repeat move/[apply].
+Qed.
+
+Lemma E_Proj2 n Γ (a b : PTm n) A B :
+  Γ ⊢ a ≡ b ∈ PBind PSig A B ->
+  Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B.
+Proof.
+  move => h.
+  have [i] : exists i, Γ ⊢ PBind PSig A B ∈ PUniv i by hauto l:on use:regularity.
+  move : E_Proj2 h; by repeat move/[apply].
+Qed.
diff --git a/theories/algorithmic.v b/theories/algorithmic.v
new file mode 100644
index 0000000..14092ef
--- /dev/null
+++ b/theories/algorithmic.v
@@ -0,0 +1,105 @@
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing.
+From Hammer Require Import Tactics.
+Require Import ssreflect ssrbool.
+Require Import Psatz.
+From stdpp Require Import relations (rtc(..)).
+Require Import Coq.Logic.FunctionalExtensionality.
+
+Module HRed.
+  Inductive R {n} : PTm n -> PTm n ->  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)
+
+  (*************** 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).
+End HRed.
+
+(* Coquand's algorithm with subtyping *)
+Reserved Notation "a ⇔ b" (at level 70).
+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 :=
+| CE_AbsAbs a b :
+  a ↔ b ->
+  (* --------------------- *)
+  PAbs a ⇔ PAbs b
+
+| CE_AbsNeu a u :
+  ishne u ->
+  a ↔ PApp (ren_PTm shift u) (VarPTm var_zero) ->
+  (* --------------------- *)
+  PAbs a ⇔ u
+
+| CE_NeuAbs a u :
+  ishne u ->
+  PApp (ren_PTm shift u) (VarPTm var_zero) ↔ a ->
+  (* --------------------- *)
+  u ⇔ PAbs a
+
+| CE_PairPair a0 a1 b0 b1 :
+  a0 ↔ a1 ->
+  b0 ↔ b1 ->
+  (* ---------------------------- *)
+  PPair a0 b0 ⇔ PPair a1 b1
+
+| CE_PairNeu a0 a1 u :
+  ishne u ->
+  a0 ↔ PProj PL u ->
+  a1 ↔ PProj PR u ->
+  (* ----------------------- *)
+  PPair a0 a1 ⇔ u
+
+| CE_NeuPair a0 a1 u :
+  ishne u ->
+  PProj PL u ↔ a0 ->
+  PProj PR u ↔ a1 ->
+  (* ----------------------- *)
+  u ⇔ PPair a0 a1
+
+| CE_ProjCong p u0 u1 :
+  ishne u0 ->
+  ishne u1 ->
+  u0 ⇔ u1 ->
+  (* ---------------------  *)
+  PProj p u0 ⇔ PProj p u1
+
+| CE_AppCong u0 u1 a0 a1 :
+  ishne u0 ->
+  ishne u1 ->
+  u0 ⇔ u1 ->
+  a0 ↔ a1 ->
+  (* ------------------------- *)
+  PApp u0 a0 ⇔ PApp u1 a1
+
+| CE_VarCong i :
+  (* -------------------------- *)
+  VarPTm i ⇔ VarPTm i
+
+| CE_UnivCong i :
+  (* -------------------------- *)
+  PUniv i ⇔ PUniv i
+
+| CE_BindCong p A0 A1 B0 B1 :
+  A0 ↔ A1 ->
+  B0 ↔ B1 ->
+  (* ---------------------------- *)
+  PBind p A0 B0 ⇔ PBind p A1 B1
+
+with CoqEq_R {n} : PTm n -> PTm n -> Prop :=
+| CE_HRed a a' b b'  :
+  rtc HRed.R a a' ->
+  rtc HRed.R b b' ->
+  a' ⇔ b' ->
+  (* ----------------------- *)
+  a ↔ b
+where "a ⇔ b" := (CoqEq a b) and "a ↔ b" := (CoqEq_R a b).
diff --git a/theories/common.v b/theories/common.v
index 02c0b83..6ad3d44 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -45,3 +45,48 @@ Proof.
   sauto.
   congruence.
 Qed.
+
+Definition ishf {n} (a : PTm n) :=
+  match a with
+  | PPair _ _ => true
+  | PAbs _ => true
+  | PUniv _ => true
+  | PBind _ _ _ => true
+  | _ => false
+  end.
+
+Fixpoint ishne {n} (a : PTm n) :=
+  match a with
+  | VarPTm _ => true
+  | PApp a _ => ishne a
+  | PProj _ a => ishne a
+  | _ => false
+  end.
+
+Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
+
+Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
+
+Definition ispair {n} (a : PTm n) :=
+  match a with
+  | PPair _ _ => true
+  | _ => false
+  end.
+
+Definition isabs {n} (a : PTm n) :=
+  match a with
+  | PAbs _ => true
+  | _ => false
+  end.
+
+Definition ishf_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  ishf (ren_PTm ξ a) = ishf a.
+Proof. case : a => //=. Qed.
+
+Definition isabs_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  isabs (ren_PTm ξ a) = isabs a.
+Proof. case : a => //=. Qed.
+
+Definition ispair_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  ispair (ren_PTm ξ a) = ispair a.
+Proof. case : a => //=. Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index b906b6a..beed0bb 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -166,41 +166,6 @@ with TRedSN {n} : PTm n -> PTm n -> Prop :=
 
 Derive Dependent Inversion tred_inv with (forall n (a b : PTm n), TRedSN a b) Sort Prop.
 
-Definition ishf {n} (a : PTm n) :=
-  match a with
-  | PPair _ _ => true
-  | PAbs _ => true
-  | PUniv _ => true
-  | PBind _ _ _ => true
-  | _ => false
-  end.
-Definition isbind {n} (a : PTm n) := if a is PBind _ _ _ then true else false.
-
-Definition isuniv {n} (a : PTm n) := if a is PUniv _ then true else false.
-
-Definition ispair {n} (a : PTm n) :=
-  match a with
-  | PPair _ _ => true
-  | _ => false
-  end.
-
-Definition isabs {n} (a : PTm n) :=
-  match a with
-  | PAbs _ => true
-  | _ => false
-  end.
-
-Definition ishf_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
-  ishf (ren_PTm ξ a) = ishf a.
-Proof. case : a => //=. Qed.
-
-Definition isabs_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
-  isabs (ren_PTm ξ a) = isabs a.
-Proof. case : a => //=. Qed.
-
-Definition ispair_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
-  ispair (ren_PTm ξ a) = ispair a.
-Proof. case : a => //=. Qed.
 
 Lemma PProj_imp n p a :
   @ishf n a ->

From c5de86339f6120cd51d0a7e53fe23565eb4d06a0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Feb 2025 21:50:23 -0500
Subject: [PATCH 078/210] Finish subject reduction

---
 theories/algorithmic.v  |  14 ++++
 theories/preservation.v | 176 ++++++++++++++++++++++++++++++++++++++++
 theories/structural.v   |  94 +++++++++++----------
 3 files changed, 243 insertions(+), 41 deletions(-)
 create mode 100644 theories/preservation.v

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 14092ef..285402d 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -103,3 +103,17 @@ with CoqEq_R {n} : PTm n -> PTm n -> Prop :=
   (* ----------------------- *)
   a ↔ b
 where "a ⇔ b" := (CoqEq a b) and "a ↔ b" := (CoqEq_R a b).
+
+Scheme coqeq_ind := Induction for CoqEq Sort Prop
+  with coqeq_r_ind := Induction for CoqEq_R Sort Prop.
+
+Combined Scheme coqeq_mutual from coqeq_ind, coqeq_r_ind.
+
+Lemma coqeq_sound_mutual : forall n,
+    (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).
+Proof.
+  apply coqeq_mutual.
+  - move => n a b ha iha Γ U wta wtb.
+    (* Need to use the fundamental lemma to show that U normalizes to a Pi type *)
+Admitted.
diff --git a/theories/preservation.v b/theories/preservation.v
new file mode 100644
index 0000000..9b7205a
--- /dev/null
+++ b/theories/preservation.v
@@ -0,0 +1,176 @@
+Require Import Autosubst2.core Autosubst2.fintype 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 :
+  Γ ⊢ 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 A B hb _ ha _ b0 a0 [*]. subst.
+    exists A,B.
+    repeat split => //=.
+    have [i] : exists i, Γ ⊢ PBind PPi A B ∈  PUniv i by sfirstorder use:regularity.
+    hauto lq:on use:bind_inst, E_Refl.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+Lemma Abs_Inv n Γ (a : PTm (S n)) U :
+  Γ ⊢ PAbs a ∈ U ->
+  exists A B, funcomp (ren_PTm shift) (scons 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 A B i hP _ ha _ a0 [*]. subst.
+    exists A, B. repeat split => //=.
+    hauto lq:on use:E_Refl, Su_Eq.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+Lemma Proj1_Inv n Γ (a : PTm n) 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 A B ha _ a0 [*]. subst.
+    exists A, B. split => //=.
+    eapply regularity in ha.
+    move : ha => [i].
+    move /Bind_Inv => [j][h _].
+    by move /E_Refl /Su_Eq in h.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+Lemma Proj2_Inv n Γ (a : PTm n) 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 A B ha _ a0 [*]. subst.
+    exists A, B. split => //=.
+    have ha' := ha.
+    eapply regularity in ha.
+    move : ha => [i ha].
+    move /T_Proj1 in ha'.
+    apply : bind_inst; eauto.
+    apply : E_Refl ha'.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+Lemma Pair_Inv n Γ (a b : PTm n) 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 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 regularity_sub0 : forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> exists i, Γ ⊢ A ∈ PUniv i.
+Proof. hauto lq:on use:regularity. Qed.
+
+Lemma E_AppAbs : forall n (a : PTm (S n)) (b : PTm n) (Γ : fin n -> PTm n) (A : PTm n),
+  Γ ⊢ PApp (PAbs a) b ∈ A -> Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ A.
+Proof.
+  move => n a b Γ A ha.
+  move /App_Inv : ha.
+  move => [A0][B0][ha][hb]hS.
+  move /Abs_Inv : ha => [A1][B1][ha]hS0.
+  have hb' := hb.
+  move /E_Refl in hb.
+  have hS1 : Γ ⊢ A0 ≲ A1 by sfirstorder use:Su_Pi_Proj1.
+  have [i hPi] : exists i, Γ ⊢ PBind PPi A1 B1 ∈ PUniv i by sfirstorder use:regularity_sub0.
+  move : Su_Pi_Proj2 hS0 hb; repeat move/[apply].
+  move : hS => /[swap]. move : Su_Transitive. repeat move/[apply].
+  move => h.
+  apply : E_Conv; eauto.
+  apply : E_AppAbs; eauto.
+  eauto using T_Conv.
+Qed.
+
+Lemma E_ProjPair1 : forall n (a b : PTm n) (Γ : fin n -> PTm n) (A : PTm n),
+    Γ ⊢ PProj PL (PPair a b) ∈ A -> Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A.
+Proof.
+  move => n 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.
+  move /Su_Sig_Proj1 in hS.
+  have {hA0} {}hS : Γ ⊢ A1 ≲ A by eauto using Su_Transitive.
+  apply : E_Conv; eauto.
+  apply : E_ProjPair1; eauto.
+Qed.
+
+Lemma RRed_Eq n Γ (a b : PTm n) A :
+  Γ ⊢ a ∈ A ->
+  RRed.R a b ->
+  Γ ⊢ a ≡ b ∈ A.
+Proof.
+  move => + h. move : Γ A. elim : n a b /h => n.
+  - apply E_AppAbs.
+  - move => p a b Γ A.
+    case : p => //=.
+    + apply E_ProjPair1.
+    + move /Proj2_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.
+      have : Γ ⊢ PPair a b ∈ PBind PSig A1 B1 by hauto lq:on ctrs:Wt.
+      move /T_Proj1.
+      move /E_ProjPair1 /E_Symmetric => h.
+      have /Su_Sig_Proj1 hSA := hS.
+      have : Γ ⊢ subst_PTm (scons a VarPTm) B1 ≲ subst_PTm (scons (PProj PL (PPair a b)) VarPTm) B0 by
+        apply : Su_Sig_Proj2; eauto.
+      move : hA0 => /[swap]. move : Su_Transitive. repeat move/[apply].
+      move {hS}.
+      move => ?. apply : E_Conv; eauto. apply : E_ProjPair2; eauto.
+  - qauto l:on use:Abs_Inv, E_Conv, regularity_sub0, E_Abs.
+  - move => a0 a1 b ha iha Γ A /App_Inv [A0][B0][ih0][ih1]hU.
+    have {}/iha iha := ih0.
+    have [i hP] : exists i, Γ ⊢ PBind PPi A0 B0 ∈ PUniv i by sfirstorder use:regularity.
+    apply : E_Conv; eauto.
+    apply : E_App; eauto using E_Refl.
+  - move => a0 b0 b1 ha iha Γ A /App_Inv [A0][B0][ih0][ih1]hU.
+    have {}/iha iha := ih1.
+    have [i hP] : exists i, Γ ⊢ PBind PPi A0 B0 ∈ PUniv i by sfirstorder use:regularity.
+    apply : E_Conv; eauto.
+    apply : E_App; eauto.
+    sfirstorder use:E_Refl.
+  - move => a0 a1 b ha iha Γ A /Pair_Inv.
+    move => [A0][B0][h0][h1]hU.
+    have [i hP] : exists i, Γ ⊢ PBind PSig A0 B0 ∈ PUniv i by eauto using regularity_sub0.
+    have {}/iha iha := h0.
+    apply : E_Conv; eauto.
+    apply : E_Pair; eauto using E_Refl.
+  - move => a b0 b1 ha iha Γ A /Pair_Inv.
+    move => [A0][B0][h0][h1]hU.
+    have [i hP] : exists i, Γ ⊢ PBind PSig A0 B0 ∈ PUniv i by eauto using regularity_sub0.
+    have {}/iha iha := h1.
+    apply : E_Conv; eauto.
+    apply : E_Pair; eauto using E_Refl.
+  - case.
+    + move => a0 a1 ha iha Γ A /Proj1_Inv [A0][B0][h0]hU.
+      apply : E_Conv; eauto.
+      qauto l:on ctrs:Eq,Wt.
+    + move => a0 a1 ha iha Γ A /Proj2_Inv [A0][B0][h0]hU.
+      have [i hP] : exists i, Γ ⊢ PBind PSig A0 B0 ∈ PUniv i by sfirstorder use:regularity.
+      apply : E_Conv; eauto.
+      apply : E_Proj2; eauto.
+  - move => p A0 A1 B hA ihA Γ U /Bind_Inv [i][h0][h1]hU.
+    have {}/ihA ihA := h0.
+    apply : E_Conv; eauto.
+    apply E_Bind'; eauto using E_Refl.
+  - move => p A0 A1 B hA ihA Γ U /Bind_Inv [i][h0][h1]hU.
+    have {}/ihA ihA := h1.
+    apply : E_Conv; eauto.
+    apply E_Bind'; eauto using E_Refl.
+Qed.
diff --git a/theories/structural.v b/theories/structural.v
index 8a6348e..84e70f3 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -36,32 +36,44 @@ Proof.
   hauto lq:on ctrs:Wt use:wff_mutual.
 Qed.
 
-
-Lemma Pi_Inv n Γ (A : PTm n) B U :
-  Γ ⊢ PBind PPi A B ∈ U ->
+Lemma Bind_Inv n Γ p (A : PTm n) B U :
+  Γ ⊢ PBind p 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.
+  move E :(PBind p A B) => T h.
+  move : p 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 Sig_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 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 :
   U = subst_PTm (scons a VarPTm) B ->
@@ -166,7 +178,7 @@ Proof.
   - hauto lq:on rew:off ctrs:Wt, Wff  use:renaming_up.
   - move => n Γ a A B i hP ihP ha iha m Δ ξ hΔ hξ.
     apply : T_Abs; eauto.
-    move : ihP(hΔ) (hξ); repeat move/[apply]. move/Pi_Inv.
+    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Δ.
@@ -177,7 +189,7 @@ Proof.
   - 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 : ihPi (hΔ) (hξ). repeat move/[apply].
-    move => /Pi_Inv [j][h0][h1]h2.
+    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.
@@ -190,14 +202,14 @@ Proof.
   - qauto l:on ctrs:Eq, LEq.
   - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ξ hΔ hξ.
     move : ihP (hξ) (hΔ). repeat move/[apply].
-    move /Pi_Inv.
+    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 : {hP} ihP (hξ) (hΔ). repeat move/[apply].
-    move /Sig_Inv => [i0][h0][h1]h2.
+    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.
@@ -317,11 +329,11 @@ Proof.
   - hauto lq:on use:morphing_up, Wff_Cons', T_Bind.
   - move => n Γ a A B i hP ihP ha iha m Δ ρ hΔ hρ.
     move : ihP (hΔ) (hρ); repeat move/[apply].
-    move /Pi_Inv => [j][h0][h1]h2. move {hP}.
+    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.
     apply : T_Abs; eauto.
     apply : iha.
-    hauto lq:on use:Wff_Cons', Pi_Inv.
+    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Δ.
@@ -336,7 +348,7 @@ Proof.
   - 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 : ihPi (hΔ) (hρ). repeat move/[apply].
-    move => /Pi_Inv [j][h0][h1]h2.
+    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.
@@ -350,14 +362,14 @@ Proof.
   - qauto l:on ctrs:Eq, LEq.
   - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ρ hΔ hρ.
     move : ihP (hρ) (hΔ). repeat move/[apply].
-    move /Pi_Inv.
+    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 : {hP} ihP (hρ) (hΔ). repeat move/[apply].
-    move /Sig_Inv => [i0][h0][h1]h2.
+    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.
@@ -484,12 +496,12 @@ Proof.
     econstructor; eauto.
     apply : renaming_shift; eauto.
   - move => n Γ b a A B hb [i ihb] ha [j iha].
-    move /Pi_Inv : ihb => [k][h0][h1]h2.
+    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:Sig_Inv.
-  - move => n Γ a A B ha [i /Sig_Inv[j][h0][h1]h2].
+  - hauto lq:on use:Bind_Inv.
+  - move => n Γ 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.
@@ -513,23 +525,23 @@ Proof.
     qauto use:T_App.
     move /E_Symmetric in ha.
     by eauto using bind_inst.
-    hauto lq:on ctrs:Wt,Eq,LEq lq:on use:Pi_Inv, substing_wt.
+    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:Sig_Inv db:wt.
+  - hauto lq:on use:Bind_Inv db:wt.
   - move => n Γ 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:Sig_Inv, substing_wt.
+    hauto lq:on ctrs:Wt,Eq,LEq use:Bind_Inv, substing_wt.
   - hauto lq:on ctrs:Wt.
   - hauto q:on use:substing_wt db:wt.
   - hauto l:on use:bind_inst db:wt.
   - move => n Γ b A B i hΓ ihΓ 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)
-                        by hauto lq:on use:weakening_wt, Pi_Inv.
+                        by hauto lq:on use:weakening_wt, Bind_Inv.
     apply : T_Abs; eauto.
     apply : T_App'; eauto; rewrite-/ren_PTm.
     by asimpl.
@@ -554,22 +566,22 @@ Proof.
     sfirstorder use:ctx_eq_subst_one.
   - sfirstorder.
   - move => n Γ A0 A1 B0 B1 _ [i][ih0 ih1].
-    move /Pi_Inv : ih0 => [i0][h _].
-    move /Pi_Inv : ih1 => [i1][h' _].
+    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 /Sig_Inv : ih0 => [i0][h _].
-    move /Sig_Inv : ih1 => [i1][h' _].
+    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 => [i][ihP0]ihP1.
     move => ha [iha0][iha1][j]ihA1.
-    move /Pi_Inv :ihP0 => [i0][ih0][ih0' _].
-    move /Pi_Inv :ihP1 => [i1][ih1][ih1' _].
+    move /Bind_Inv :ihP0 => [i0][ih0][ih0' _].
+    move /Bind_Inv :ihP1 => [i1][ih1][ih1' _].
     have [*] : i0 <= max i0 i1 /\ i1 <= max i0 i1 by lia.
     exists (max i0 i1).
     split.
@@ -581,8 +593,8 @@ Proof.
   - move => n Γ a0 a1 A0 A1 B0 B1 /Su_Sig_Proj1 hA1.
     move => [i][ihP0]ihP1.
     move => ha [iha0][iha1][j]ihA1.
-    move /Sig_Inv :ihP0 => [i0][ih0][ih0' _].
-    move /Sig_Inv :ihP1 => [i1][ih1][ih1' _].
+    move /Bind_Inv :ihP0 => [i0][ih0][ih0' _].
+    move /Bind_Inv :ihP1 => [i1][ih1][ih1' _].
     have [*] : i0 <= max i0 i1 /\ i1 <= max i0 i1 by lia.
     exists (max i0 i1).
     split.

From c1a8e9d2e1669fa8fe7a62c704ba59f67eb1910f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Feb 2025 21:51:27 -0500
Subject: [PATCH 079/210] Add the top-level subject reduction theorem

---
 theories/preservation.v | 6 ++++++
 1 file changed, 6 insertions(+)

diff --git a/theories/preservation.v b/theories/preservation.v
index 9b7205a..d1f9a3e 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -174,3 +174,9 @@ Proof.
     apply : E_Conv; eauto.
     apply E_Bind'; eauto using E_Refl.
 Qed.
+
+Theorem subject_reduction n Γ (a b A : PTm n) :
+  Γ ⊢ a ∈ A ->
+  RRed.R a b ->
+  Γ ⊢ b ∈ A.
+Proof. hauto lq:on use:RRed_Eq, regularity. Qed.

From 15f8a9c6878f114abd85950c3b1220a88c7d8bb2 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 11 Feb 2025 19:15:06 -0500
Subject: [PATCH 080/210] Try a few cases of soundness

---
 theories/algorithmic.v | 204 ++++++++++++++++++++++++++++++++---------
 1 file changed, 160 insertions(+), 44 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 285402d..dc3c083 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing.
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing preservation admissible.
 From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
 Require Import Psatz.
@@ -24,96 +24,212 @@ Module HRed.
 End HRed.
 
 (* Coquand's algorithm with subtyping *)
-Reserved Notation "a ⇔ b" (at level 70).
+Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
+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 :=
 | CE_AbsAbs a b :
-  a ↔ b ->
+  a ⇔ b ->
   (* --------------------- *)
-  PAbs a ⇔ PAbs b
+  PAbs a ↔ PAbs b
 
 | CE_AbsNeu a u :
   ishne u ->
-  a ↔ PApp (ren_PTm shift u) (VarPTm var_zero) ->
+  a ⇔ PApp (ren_PTm shift u) (VarPTm var_zero) ->
   (* --------------------- *)
-  PAbs a ⇔ u
+  PAbs a ↔ u
 
 | CE_NeuAbs a u :
   ishne u ->
-  PApp (ren_PTm shift u) (VarPTm var_zero) ↔ a ->
+  PApp (ren_PTm shift u) (VarPTm var_zero) ⇔ a ->
   (* --------------------- *)
-  u ⇔ PAbs a
+  u ↔ PAbs a
 
 | CE_PairPair a0 a1 b0 b1 :
-  a0 ↔ a1 ->
-  b0 ↔ b1 ->
+  a0 ⇔ a1 ->
+  b0 ⇔ b1 ->
   (* ---------------------------- *)
-  PPair a0 b0 ⇔ PPair a1 b1
+  PPair a0 b0 ↔ PPair a1 b1
 
 | CE_PairNeu a0 a1 u :
   ishne u ->
-  a0 ↔ PProj PL u ->
-  a1 ↔ PProj PR u ->
+  a0 ⇔ PProj PL u ->
+  a1 ⇔ PProj PR u ->
   (* ----------------------- *)
-  PPair a0 a1 ⇔ u
+  PPair a0 a1 ↔ u
 
 | CE_NeuPair a0 a1 u :
   ishne u ->
-  PProj PL u ↔ a0 ->
-  PProj PR u ↔ a1 ->
+  PProj PL u ⇔ a0 ->
+  PProj PR u ⇔ a1 ->
   (* ----------------------- *)
-  u ⇔ PPair a0 a1
+  u ↔ PPair a0 a1
+
+| CE_UnivCong i :
+  (* -------------------------- *)
+  PUniv i ↔ PUniv i
+
+| CE_BindCong p A0 A1 B0 B1 :
+  A0 ⇔ A1 ->
+  B0 ⇔ B1 ->
+  (* ---------------------------- *)
+  PBind p A0 B0 ↔ PBind p A1 B1
+
+| CE_NeuNeu a0 a1 :
+  a0 ∼ a1 ->
+  a0 ↔ a1
+
+with CoqEq_Neu  {n} : PTm n -> PTm n -> Prop :=
+| CE_VarCong i :
+  (* -------------------------- *)
+  VarPTm i ∼ VarPTm i
 
 | CE_ProjCong p u0 u1 :
   ishne u0 ->
   ishne u1 ->
-  u0 ⇔ u1 ->
+  u0 ∼ u1 ->
   (* ---------------------  *)
-  PProj p u0 ⇔ PProj p u1
+  PProj p u0 ∼ PProj p u1
 
 | CE_AppCong u0 u1 a0 a1 :
   ishne u0 ->
   ishne u1 ->
-  u0 ⇔ u1 ->
-  a0 ↔ a1 ->
+  u0 ∼ u1 ->
+  a0 ⇔ a1 ->
   (* ------------------------- *)
-  PApp u0 a0 ⇔ PApp u1 a1
-
-| CE_VarCong i :
-  (* -------------------------- *)
-  VarPTm i ⇔ VarPTm i
-
-| CE_UnivCong i :
-  (* -------------------------- *)
-  PUniv i ⇔ PUniv i
-
-| CE_BindCong p A0 A1 B0 B1 :
-  A0 ↔ A1 ->
-  B0 ↔ B1 ->
-  (* ---------------------------- *)
-  PBind p A0 B0 ⇔ PBind p A1 B1
+  PApp u0 a0 ∼ PApp u1 a1
 
 with CoqEq_R {n} : PTm n -> PTm n -> Prop :=
 | CE_HRed a a' b b'  :
   rtc HRed.R a a' ->
   rtc HRed.R b b' ->
-  a' ⇔ b' ->
+  a' ↔ b' ->
   (* ----------------------- *)
-  a ↔ b
-where "a ⇔ b" := (CoqEq a b) and "a ↔ b" := (CoqEq_R a b).
+  a ⇔ b
+where "a ↔ b" := (CoqEq a b) and "a ⇔ b" := (CoqEq_R a b) and "a ∼ b" := (CoqEq_Neu a b).
 
-Scheme coqeq_ind := Induction for CoqEq Sort Prop
+Scheme
+  coqeq_neu_ind := Induction for CoqEq_Neu Sort Prop
+  with coqeq_ind := Induction for CoqEq Sort Prop
   with coqeq_r_ind := Induction for CoqEq_R Sort Prop.
 
-Combined Scheme coqeq_mutual from coqeq_ind, coqeq_r_ind.
+Combined Scheme coqeq_mutual from coqeq_neu_ind, coqeq_ind, coqeq_r_ind.
+
+Lemma Var_Inv n Γ (i : fin n) A :
+  Γ ⊢ VarPTm i ∈ A ->
+  ⊢ Γ /\ Γ ⊢ Γ i ≲ A.
+Proof.
+  move E : (VarPTm i) => u hu.
+  move : i E.
+  elim : n Γ u A / hu=>//=.
+  - move => n Γ i hΓ i0 [?]. subst.
+    repeat split => //=.
+    have h : Γ ⊢ VarPTm i ∈ Γ i by eauto using T_Var.
+    eapply structural.regularity in h.
+    move : h => [i0]?.
+    apply : Su_Eq. apply E_Refl; eassumption.
+  - sfirstorder use:Su_Transitive.
+Qed.
 
 Lemma coqeq_sound_mutual : forall n,
-    (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 n), 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).
 Proof.
   apply coqeq_mutual.
-  - move => n a b ha iha Γ U wta wtb.
+  - move => n i Γ A B hi0 hi1.
+    move /Var_Inv : hi0 => [hΓ h0].
+    move /Var_Inv : hi1 => [_ h1].
+    exists (Γ i).
+    repeat split => //=.
+    apply E_Refl. eauto using T_Var.
+  - move => n [] u0 u1 hu0 hu1 hu ihu Γ A B hu0' hu1'.
+    + move /Proj1_Inv : hu0'.
+      move => [A0][B0][hu0']hu0''.
+      move /Proj1_Inv : hu1'.
+      move => [A1][B1][hu1']hu1''.
+      specialize ihu with (1 := hu0') (2 := hu1').
+      move : ihu.
+      move => [C][ih0][ih1]ih.
+      have [A2 [B2 [{}ih0 [{}ih1 {}ih]]]] : exists A2 B2, Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 /\ Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by admit.
+
+      have /Su_Sig_Proj1 hs0 := ih0.
+      have /Su_Sig_Proj1 hs1 := ih1.
+      exists A2.
+      repeat split; eauto using Su_Transitive.
+      apply : E_Proj1; eauto.
+    + move /Proj2_Inv : hu0'.
+      move => [A0][B0][hu0']hu0''.
+      move /Proj2_Inv : hu1'.
+      move => [A1][B1][hu1']hu1''.
+      specialize ihu with (1 := hu0') (2 := hu1').
+      move : ihu.
+      move => [C][ih0][ih1]ih.
+      have [A2 [B2 [{}ih0 [{}ih1 {}ih]]]] : exists A2 B2, Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 /\ Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by admit.
+      exists (subst_PTm (scons (PProj PL u0) VarPTm) B2).
+      have [? ?] : Γ ⊢ u0 ∈ PBind PSig A2 B2 /\ Γ ⊢ u1 ∈ PBind PSig A2 B2 by hauto l:on use:structural.regularity.
+      repeat split => //=.
+      apply : Su_Transitive ;eauto.
+      apply : Su_Sig_Proj2; eauto.
+      apply E_Refl. eauto using T_Proj1.
+      apply : Su_Transitive ;eauto.
+      apply : Su_Sig_Proj2; eauto.
+      apply : E_Proj1; eauto.
+      move /regularity_sub0 : ih1 => [i ?].
+      apply : E_Proj2; eauto.
+  - move => n u0 u1 a0 a1 neu0 neu1 hu ihu ha iha Γ A B wta0 wta1.
+    admit.
+  - move => n 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 admit.
+    have ? : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
+    have ? : Γ ⊢ PBind PPi A1 B1 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
+    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.
+    eauto using E_Symmetric, Su_Eq.
+    apply : E_Abs; eauto. hauto l:on use:structural.regularity.
+    apply iha.
+    eapply structural.ctx_eq_subst_one with (A0 := A0); eauto.
+    admit.
+    admit.
+    admit.
+    eapply structural.ctx_eq_subst_one with (A0 := A1); eauto.
+    admit.
+    admit.
+    admit.
     (* Need to use the fundamental lemma to show that U normalizes to a Pi type *)
+  - move => n 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 admit.
+    have hPi'' : Γ ⊢ PBind PPi A2 B2 ≲ A by eauto using Su_Eq, Su_Transitive, E_Symmetric.
+    have hPi' : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Eq, Su_Transitive.
+    apply E_Conv with (A := PBind PPi A2 B2); eauto.
+    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:structural.wff_mutual.
+    hauto l:on use:structural.regularity.
+    apply T_Conv with (A := A);eauto.
+    eauto using Su_Eq.
+    move => ?.
+    suff : Γ ⊢ PAbs a ≡ PAbs (PApp (ren_PTm shift u) (VarPTm var_zero)) ∈ PBind PPi A2 B2
+      by eauto using E_Transitive.
+    apply : E_Abs; eauto. hauto l:on use:structural.regularity.
+    apply iha.
+    admit.
+    admit.
+  (* Mirrors the last case *)
+  - admit.
+  - admit.
+  - admit.
+  - admit.
+  - sfirstorder use:E_Refl.
+  - admit.
+  - hauto lq:on ctrs:Eq,LEq,Wt.
+  - move => n a a' b b' ha hb.
+    admit.
 Admitted.

From 823f61d89fc5ea8683fe11b1fcd36a6f281d21f2 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 12 Feb 2025 15:54:42 -0500
Subject: [PATCH 081/210] Finish most cases of the soundness proof

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

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index dc3c083..fdcd1ab 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1,4 +1,5 @@
-Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax common typing preservation admissible.
+Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
+  common typing preservation admissible fp_red.
 From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
 Require Import Psatz.
@@ -21,8 +22,30 @@ Module HRed.
   | ProjCong p a0 a1 :
     R a0 a1 ->
     R (PProj p a0) (PProj p a1).
+
+  Lemma ToRRed n (a b : PTm n) : 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.
+  Proof.
+    sfirstorder use:subject_reduction, ToRRed.
+  Qed.
+
+  Lemma ToEq n Γ (a b : PTm n) 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.
+  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.
+  Proof.
+    induction 2; sauto lq:on use:HRed.ToEq, E_Transitive, HRed.preservation.
+  Qed.
+End HReds.
+
 (* Coquand's algorithm with subtyping *)
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
@@ -181,6 +204,15 @@ Proof.
       move /regularity_sub0 : ih1 => [i ?].
       apply : E_Proj2; eauto.
   - move => n 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].
+    move => [C][hC0][hC1]hu01.
+    have [i [A2 [B2 hPi]]] : exists i A2 B2, Γ ⊢ PBind PPi A2 B2 ≡ C ∈ PUniv i by admit.
+    exists (subst_PTm (scons a0 VarPTm) B2).
+    repeat split. admit. admit.
+    apply E_App with (A := A2). apply : E_Conv; eauto.
+    apply : Su_Eq; apply : E_Symmetric; eauto.
     admit.
   - move => n a b ha iha Γ A h0 h1.
     move /Abs_Inv : h0 => [A0][B0][h0]h0'.
@@ -221,15 +253,47 @@ Proof.
     apply : E_Abs; eauto. hauto l:on use:structural.regularity.
     apply iha.
     admit.
+    eapply structural.T_App' with (A := ren_PTm shift A2) (B := ren_PTm (upRen_PTm_PTm shift) B2). by asimpl.
+    eapply structural.weakening_wt' with (a := u) (A := PBind PPi A2 B2). reflexivity. by asimpl.
+    admit.
+    apply : T_Conv; eauto. apply : Su_Eq; eauto.
+    apply T_Var. apply : structural.Wff_Cons'; eauto.
     admit.
   (* Mirrors the last case *)
   - admit.
-  - admit.
-  - admit.
+  - move => n 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 admit.
+    apply E_Conv with (A := PBind PSig A2 B2); last by eauto using E_Symmetric, Su_Eq.
+    apply : E_Pair; eauto. hauto l:on use:structural.regularity.
+    apply iha. admit. admit.
+    admit.
+  - move => n 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 admit.
+    have hA' : Γ ⊢ PBind PSig A2 B2 ≲ A by eauto using E_Symmetric, Su_Eq.
+    move /E_Conv : (hA'). apply.
+    apply : E_Transitive; last (apply E_Symmetric; apply : E_PairEta).
+    apply : E_Pair; eauto. hauto l:on use:structural.regularity.
+    apply ih0. admit. admit.
+    apply ih1. admit. admit.
+    hauto l:on use:structural.regularity.
+    apply : T_Conv; eauto using Su_Eq.
+  (* Same as before *)
   - admit.
   - sfirstorder use:E_Refl.
-  - admit.
-  - hauto lq:on ctrs:Eq,LEq,Wt.
-  - move => n a a' b b' ha hb.
+  - move => n p A0 A1 B0 B1 hA ihA hB ihB Γ A hA0 hA1.
+    move /structural.Bind_Inv : hA0 => [i][hA0][hB0]hU.
+    move /structural.Bind_Inv : hA1 => [j][hA1][hB1]hU1.
+    have [l [k hk]] : exists l k, Γ ⊢ A ≡ PUniv k ∈ PUniv l by admit.
+    apply E_Conv with (A := PUniv k); last by eauto using Su_Eq, E_Symmetric.
+    apply E_Bind. admit.
+    apply ihB.
     admit.
+    admit.
+  - hauto lq:on ctrs:Eq,LEq,Wt.
+  - move => n 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.
 Admitted.

From 5ac2bf1c400d51a478d0b829e8e0c66b9da3b04c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 12 Feb 2025 15:56:35 -0500
Subject: [PATCH 082/210] Minor

---
 theories/algorithmic.v | 48 ++++++++++++++----------------------------
 theories/structural.v  | 16 ++++++++++++++
 2 files changed, 32 insertions(+), 32 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index fdcd1ab..28e194b 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1,5 +1,5 @@
 Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
-  common typing preservation admissible fp_red.
+  common typing preservation admissible fp_red structural.
 From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
 Require Import Psatz.
@@ -140,22 +140,6 @@ Scheme
 
 Combined Scheme coqeq_mutual from coqeq_neu_ind, coqeq_ind, coqeq_r_ind.
 
-Lemma Var_Inv n Γ (i : fin n) A :
-  Γ ⊢ VarPTm i ∈ A ->
-  ⊢ Γ /\ Γ ⊢ Γ i ≲ A.
-Proof.
-  move E : (VarPTm i) => u hu.
-  move : i E.
-  elim : n Γ u A / hu=>//=.
-  - move => n Γ i hΓ i0 [?]. subst.
-    repeat split => //=.
-    have h : Γ ⊢ VarPTm i ∈ Γ i by eauto using T_Var.
-    eapply structural.regularity in h.
-    move : h => [i0]?.
-    apply : Su_Eq. apply E_Refl; eassumption.
-  - sfirstorder use:Su_Transitive.
-Qed.
-
 Lemma coqeq_sound_mutual : forall n,
     (forall (a b : PTm n), a ∼ b -> forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> exists C,
        Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ≡ b ∈ C) /\
@@ -193,7 +177,7 @@ Proof.
       move => [C][ih0][ih1]ih.
       have [A2 [B2 [{}ih0 [{}ih1 {}ih]]]] : exists A2 B2, Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 /\ Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by admit.
       exists (subst_PTm (scons (PProj PL u0) VarPTm) B2).
-      have [? ?] : Γ ⊢ u0 ∈ PBind PSig A2 B2 /\ Γ ⊢ u1 ∈ PBind PSig A2 B2 by hauto l:on use:structural.regularity.
+      have [? ?] : Γ ⊢ u0 ∈ PBind PSig A2 B2 /\ Γ ⊢ u1 ∈ PBind PSig A2 B2 by hauto l:on use:regularity.
       repeat split => //=.
       apply : Su_Transitive ;eauto.
       apply : Su_Sig_Proj2; eauto.
@@ -223,13 +207,13 @@ Proof.
     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.
     eauto using E_Symmetric, Su_Eq.
-    apply : E_Abs; eauto. hauto l:on use:structural.regularity.
+    apply : E_Abs; eauto. hauto l:on use:regularity.
     apply iha.
-    eapply structural.ctx_eq_subst_one with (A0 := A0); eauto.
+    eapply ctx_eq_subst_one with (A0 := A0); eauto.
     admit.
     admit.
     admit.
-    eapply structural.ctx_eq_subst_one with (A0 := A1); eauto.
+    eapply ctx_eq_subst_one with (A0 := A1); eauto.
     admit.
     admit.
     admit.
@@ -243,21 +227,21 @@ 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:structural.wff_mutual.
-    hauto l:on use:structural.regularity.
+    sfirstorder use:wff_mutual.
+    hauto l:on use:regularity.
     apply T_Conv with (A := A);eauto.
     eauto using Su_Eq.
     move => ?.
     suff : Γ ⊢ PAbs a ≡ PAbs (PApp (ren_PTm shift u) (VarPTm var_zero)) ∈ PBind PPi A2 B2
       by eauto using E_Transitive.
-    apply : E_Abs; eauto. hauto l:on use:structural.regularity.
+    apply : E_Abs; eauto. hauto l:on use:regularity.
     apply iha.
     admit.
-    eapply structural.T_App' with (A := ren_PTm shift A2) (B := ren_PTm (upRen_PTm_PTm shift) B2). by asimpl.
-    eapply structural.weakening_wt' with (a := u) (A := PBind PPi A2 B2). reflexivity. by asimpl.
+    eapply T_App' with (A := ren_PTm shift A2) (B := ren_PTm (upRen_PTm_PTm shift) B2). by asimpl.
+    eapply weakening_wt' with (a := u) (A := PBind PPi A2 B2). reflexivity. by asimpl.
     admit.
     apply : T_Conv; eauto. apply : Su_Eq; eauto.
-    apply T_Var. apply : structural.Wff_Cons'; eauto.
+    apply T_Var. apply : Wff_Cons'; eauto.
     admit.
   (* Mirrors the last case *)
   - admit.
@@ -266,7 +250,7 @@ Proof.
     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 admit.
     apply E_Conv with (A := PBind PSig A2 B2); last by eauto using E_Symmetric, Su_Eq.
-    apply : E_Pair; eauto. hauto l:on use:structural.regularity.
+    apply : E_Pair; eauto. hauto l:on use:regularity.
     apply iha. admit. admit.
     admit.
   - move => n a0 a1 u neu h0 ih0 h1 ih1 Γ A ha hu.
@@ -275,17 +259,17 @@ Proof.
     have hA' : Γ ⊢ PBind PSig A2 B2 ≲ A by eauto using E_Symmetric, Su_Eq.
     move /E_Conv : (hA'). apply.
     apply : E_Transitive; last (apply E_Symmetric; apply : E_PairEta).
-    apply : E_Pair; eauto. hauto l:on use:structural.regularity.
+    apply : E_Pair; eauto. hauto l:on use:regularity.
     apply ih0. admit. admit.
     apply ih1. admit. admit.
-    hauto l:on use:structural.regularity.
+    hauto l:on use:regularity.
     apply : T_Conv; eauto using Su_Eq.
   (* Same as before *)
   - admit.
   - sfirstorder use:E_Refl.
   - move => n p A0 A1 B0 B1 hA ihA hB ihB Γ A hA0 hA1.
-    move /structural.Bind_Inv : hA0 => [i][hA0][hB0]hU.
-    move /structural.Bind_Inv : hA1 => [j][hA1][hB1]hU1.
+    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 by admit.
     apply E_Conv with (A := PUniv k); last by eauto using Su_Eq, E_Symmetric.
     apply E_Bind. admit.
diff --git a/theories/structural.v b/theories/structural.v
index 84e70f3..9798d5b 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -604,3 +604,19 @@ Proof.
       have : Γ ⊢ a1 ∈ A1 by eauto using T_Conv.
       move : substing_wt ih1';repeat move/[apply]. by asimpl.
 Qed.
+
+Lemma Var_Inv n Γ (i : fin n) A :
+  Γ ⊢ VarPTm i ∈ A ->
+  ⊢ Γ /\ Γ ⊢ Γ i ≲ A.
+Proof.
+  move E : (VarPTm i) => u hu.
+  move : i E.
+  elim : n Γ u A / hu=>//=.
+  - move => n Γ i hΓ i0 [?]. subst.
+    repeat split => //=.
+    have h : Γ ⊢ VarPTm i ∈ Γ i by eauto using T_Var.
+    eapply regularity in h.
+    move : h => [i0]?.
+    apply : Su_Eq. apply E_Refl; eassumption.
+  - sfirstorder use:Su_Transitive.
+Qed.

From 761e96f414f320d81ec8a1ca8bee98c9b92215c2 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 12 Feb 2025 16:14:51 -0500
Subject: [PATCH 083/210] Use the abstract tactic to finish off the symmetric
 casesa

---
 theories/algorithmic.v | 40 +++++++++++++++++++++++++++++++---------
 1 file changed, 31 insertions(+), 9 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 28e194b..08873d4 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -140,12 +140,19 @@ 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).
+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,
        Γ ⊢ 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).
 Proof.
+  move => [:hAppL hPairL].
   apply coqeq_mutual.
   - move => n i Γ A B hi0 hi1.
     move /Var_Inv : hi0 => [hΓ h0].
@@ -218,7 +225,8 @@ Proof.
     admit.
     admit.
     (* Need to use the fundamental lemma to show that U normalizes to a Pi type *)
-  - move => n a u hneu ha iha Γ A wta wtu.
+  - abstract : hAppL.
+    move => n 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 admit.
     have hPi'' : Γ ⊢ PBind PPi A2 B2 ≲ A by eauto using Su_Eq, Su_Transitive, E_Symmetric.
@@ -244,8 +252,12 @@ Proof.
     apply T_Var. apply : Wff_Cons'; eauto.
     admit.
   (* Mirrors the last case *)
-  - admit.
-  - move => n a0 a1 b0 b1 ha iha hb ihb Γ A.
+  - move => n 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 /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 admit.
@@ -253,7 +265,9 @@ Proof.
     apply : E_Pair; eauto. hauto l:on use:regularity.
     apply iha. admit. admit.
     admit.
-  - move => n a0 a1 u neu h0 ih0 h1 ih1 Γ A ha hu.
+  - move => {hAppL}.
+    abstract : hPairL.
+    move => n 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 admit.
     have hA' : Γ ⊢ PBind PSig A2 B2 ≲ A by eauto using E_Symmetric, Su_Eq.
@@ -265,17 +279,25 @@ Proof.
     hauto l:on use:regularity.
     apply : T_Conv; eauto using Su_Eq.
   (* Same as before *)
-  - admit.
+  - move {hAppL}.
+    move => *. apply E_Symmetric. apply : hPairL;
+      sfirstorder use:coqeq_symmetric_mutual, E_Symmetric.
   - sfirstorder use:E_Refl.
-  - move => n p A0 A1 B0 B1 hA ihA hB ihB Γ A hA0 hA1.
+  - move => {hAppL hPairL} n 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 by admit.
     apply E_Conv with (A := PUniv k); last by eauto using Su_Eq, E_Symmetric.
-    apply E_Bind. admit.
+    move => [:eqA].
+    apply E_Bind. abstract : eqA. apply ihA.
+    apply T_Conv with (A := PUniv i); eauto.
+    by eauto using Su_Transitive, Su_Eq, E_Symmetric.
+    apply T_Conv with (A := PUniv j); eauto.
+    by eauto using Su_Transitive, Su_Eq, E_Symmetric.
     apply ihB.
-    admit.
-    admit.
+    apply T_Conv with (A := PUniv i); eauto. admit.
+    apply T_Conv with (A := PUniv j); eauto.
+    apply : ctx_eq_subst_one; eauto. apply : Su_Eq; apply eqA.  admit.
   - hauto lq:on ctrs:Eq,LEq,Wt.
   - move => n a a' b b' ha hb hab ihab Γ A ha0 hb0.
     have [*] : Γ ⊢ a' ∈ A /\ Γ ⊢ b' ∈ A by eauto using HReds.preservation.

From fa80294c5d0b2241af6bf01de65026916f35bfb3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 12 Feb 2025 16:40:51 -0500
Subject: [PATCH 084/210] Minor

---
 theories/algorithmic.v | 13 +++++++++----
 1 file changed, 9 insertions(+), 4 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 08873d4..d7d9f03 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -200,11 +200,16 @@ Proof.
     move : ihu hu0 hu1. repeat move/[apply].
     move => [C][hC0][hC1]hu01.
     have [i [A2 [B2 hPi]]] : exists i A2 B2, Γ ⊢ PBind PPi A2 B2 ≡ C ∈ PUniv i by admit.
+    have ? : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by eauto using Su_Eq, Su_Transitive.
+    have ? : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by eauto using Su_Eq, Su_Transitive.
     exists (subst_PTm (scons a0 VarPTm) B2).
-    repeat split. admit. admit.
-    apply E_App with (A := A2). apply : E_Conv; eauto.
-    apply : Su_Eq; apply : E_Symmetric; eauto.
-    admit.
+    repeat split. apply : Su_Transitive; eauto.
+    apply : Su_Pi_Proj2'; eauto using E_Refl.
+    apply : Su_Transitive; eauto.
+    eapply Su_Pi_Proj2' with (A0 := A2) (A1 := A2); eauto using E_Refl. admit.
+    apply iha. admit. admit.
+    apply E_App with (A := A2); eauto.
+    admit. admit.
   - move => n a b ha iha Γ A h0 h1.
     move /Abs_Inv : h0 => [A0][B0][h0]h0'.
     move /Abs_Inv : h1 => [A1][B1][h1]h1'.

From d053f93100b185592a4efa459101375289d279c6 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 12 Feb 2025 19:27:42 -0500
Subject: [PATCH 085/210] Finish the app neutral case

---
 theories/algorithmic.v | 67 ++++++++++++++++++++++++++++++++----------
 1 file changed, 52 insertions(+), 15 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index d7d9f03..36eb116 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -46,6 +46,37 @@ Module HReds.
   Qed.
 End HReds.
 
+Lemma T_Conv_E n Γ (a : PTm n) 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 :
+  Γ ⊢ 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 renaming_su' : forall n m Δ Γ (ξ : fin n -> fin m) (A B : PTm n) u U ,
+    u = ren_PTm ξ A ->
+    U = ren_PTm ξ B ->
+    Γ ⊢ A ≲ B ->
+    ⊢ Δ -> renaming_ok Δ Γ ξ ->
+    Δ ⊢ u ≲ U.
+Proof. move => > -> ->. hauto l:on use:renaming. Qed.
+
+Lemma weakening_su : forall n Γ (A0 A1 B : PTm n) i,
+    Γ ⊢ B ∈ PUniv i ->
+    Γ ⊢ A0 ≲ A1 ->
+    funcomp (ren_PTm shift) (scons B Γ) ⊢ ren_PTm shift A0 ≲ ren_PTm shift A1.
+Proof.
+  move => n Γ A0 A1 B i hB hlt.
+  apply : renaming_su'; eauto.
+  apply : Wff_Cons'; eauto.
+  apply : renaming_shift; eauto.
+Qed.
+
 (* Coquand's algorithm with subtyping *)
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
@@ -168,13 +199,12 @@ Proof.
       specialize ihu with (1 := hu0') (2 := hu1').
       move : ihu.
       move => [C][ih0][ih1]ih.
-      have [A2 [B2 [{}ih0 [{}ih1 {}ih]]]] : exists A2 B2, Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 /\ Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by admit.
-
-      have /Su_Sig_Proj1 hs0 := ih0.
-      have /Su_Sig_Proj1 hs1 := ih1.
-      exists A2.
-      repeat split; eauto using Su_Transitive.
-      apply : E_Proj1; eauto.
+      have [i[A2[B2 h2]]] : exists i A2 B2, Γ ⊢ PBind PSig A2 B2 ≡ C ∈ PUniv i by admit.
+      exists  A2.
+      have [h3 h4] : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 by qauto l:on use:Su_Eq, Su_Transitive.
+      repeat split;
+        eauto using Su_Sig_Proj1, Su_Transitive;[idtac].
+      apply E_Proj1 with (B := B2); eauto using E_Conv_E.
     + move /Proj2_Inv : hu0'.
       move => [A0][B0][hu0']hu0''.
       move /Proj2_Inv : hu1'.
@@ -182,7 +212,9 @@ Proof.
       specialize ihu with (1 := hu0') (2 := hu1').
       move : ihu.
       move => [C][ih0][ih1]ih.
-      have [A2 [B2 [{}ih0 [{}ih1 {}ih]]]] : exists A2 B2, Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 /\ Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by admit.
+      have [A2 [B2 [i hi]]] : exists A2 B2 i, Γ ⊢ PBind PSig A2 B2 ≡ C ∈ PUniv i by admit.
+      have [h3 h4] : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 by qauto l:on use:Su_Eq, Su_Transitive.
+      have h5 : Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by eauto using E_Conv_E.
       exists (subst_PTm (scons (PProj PL u0) VarPTm) B2).
       have [? ?] : Γ ⊢ u0 ∈ PBind PSig A2 B2 /\ Γ ⊢ u1 ∈ PBind PSig A2 B2 by hauto l:on use:regularity.
       repeat split => //=.
@@ -192,7 +224,6 @@ Proof.
       apply : Su_Transitive ;eauto.
       apply : Su_Sig_Proj2; eauto.
       apply : E_Proj1; eauto.
-      move /regularity_sub0 : ih1 => [i ?].
       apply : E_Proj2; eauto.
   - move => n u0 u1 a0 a1 neu0 neu1 hu ihu ha iha Γ A B wta0 wta1.
     move /App_Inv : wta0 => [A0][B0][hu0][ha0]hU.
@@ -201,15 +232,21 @@ Proof.
     move => [C][hC0][hC1]hu01.
     have [i [A2 [B2 hPi]]] : exists i A2 B2, Γ ⊢ PBind PPi A2 B2 ≡ C ∈ PUniv i by admit.
     have ? : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by eauto using Su_Eq, Su_Transitive.
-    have ? : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by eauto using Su_Eq, Su_Transitive.
+    have h : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by eauto using Su_Eq, Su_Transitive.
+    have ha' : Γ ⊢ a0 ≡ a1 ∈ A2 by
+      sauto lq:on use:Su_Transitive, Su_Pi_Proj1.
+    have hwf : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i by hauto l:on use:regularity.
+    have [j hj'] : exists j,Γ ⊢ A2 ∈ PUniv j by hauto l:on use:regularity.
+    have ? : ⊢ Γ by sfirstorder use:wff_mutual.
     exists (subst_PTm (scons a0 VarPTm) B2).
     repeat split. apply : Su_Transitive; eauto.
-    apply : Su_Pi_Proj2'; eauto using E_Refl.
+    apply : Su_Pi_Proj2';  eauto using E_Refl.
     apply : Su_Transitive; eauto.
-    eapply Su_Pi_Proj2' with (A0 := A2) (A1 := A2); eauto using E_Refl. admit.
-    apply iha. admit. admit.
-    apply E_App with (A := A2); eauto.
-    admit. admit.
+    have ? : Γ ⊢ A1 ≲ A2 by eauto using Su_Pi_Proj1.
+    apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2);
+      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 => n a b ha iha Γ A h0 h1.
     move /Abs_Inv : h0 => [A0][B0][h0]h0'.
     move /Abs_Inv : h1 => [A1][B1][h1]h1'.

From 48adb34946294ffb8035a332d02093037dc65b23 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 12 Feb 2025 19:53:20 -0500
Subject: [PATCH 086/210] Add simplified projection lemma

---
 theories/algorithmic.v  | 22 +++-------------------
 theories/preservation.v |  3 ---
 theories/structural.v   | 38 ++++++++++++++++++++++++++++++++++++++
 3 files changed, 41 insertions(+), 22 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 36eb116..34b60f8 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -58,25 +58,6 @@ Lemma E_Conv_E n Γ (a b : PTm n) A B i :
   Γ ⊢ a ≡ b ∈ B.
 Proof. qauto use:E_Conv, Su_Eq, E_Symmetric. Qed.
 
-Lemma renaming_su' : forall n m Δ Γ (ξ : fin n -> fin m) (A B : PTm n) u U ,
-    u = ren_PTm ξ A ->
-    U = ren_PTm ξ B ->
-    Γ ⊢ A ≲ B ->
-    ⊢ Δ -> renaming_ok Δ Γ ξ ->
-    Δ ⊢ u ≲ U.
-Proof. move => > -> ->. hauto l:on use:renaming. Qed.
-
-Lemma weakening_su : forall n Γ (A0 A1 B : PTm n) i,
-    Γ ⊢ B ∈ PUniv i ->
-    Γ ⊢ A0 ≲ A1 ->
-    funcomp (ren_PTm shift) (scons B Γ) ⊢ ren_PTm shift A0 ≲ ren_PTm shift A1.
-Proof.
-  move => n Γ A0 A1 B i hB hlt.
-  apply : renaming_su'; eauto.
-  apply : Wff_Cons'; eauto.
-  apply : renaming_shift; eauto.
-Qed.
-
 (* Coquand's algorithm with subtyping *)
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
@@ -253,6 +234,9 @@ Proof.
     have [i [A2 [B2 h]]] : exists i A2 B2, Γ ⊢ A ≡ PBind PPi A2 B2 ∈ PUniv i by admit.
     have ? : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
     have ? : Γ ⊢ 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 [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.
     eauto using E_Symmetric, Su_Eq.
diff --git a/theories/preservation.v b/theories/preservation.v
index d1f9a3e..6fe081d 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -76,9 +76,6 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-Lemma regularity_sub0 : forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> exists i, Γ ⊢ A ∈ PUniv i.
-Proof. hauto lq:on use:regularity. Qed.
-
 Lemma E_AppAbs : forall n (a : PTm (S n)) (b : PTm n) (Γ : fin n -> PTm n) (A : PTm n),
   Γ ⊢ PApp (PAbs a) b ∈ A -> Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ A.
 Proof.
diff --git a/theories/structural.v b/theories/structural.v
index 9798d5b..16fc334 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -620,3 +620,41 @@ Proof.
     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 ,
+    u = ren_PTm ξ A ->
+    U = ren_PTm ξ B ->
+    Γ ⊢ A ≲ B ->
+    ⊢ Δ -> renaming_ok Δ Γ ξ ->
+    Δ ⊢ u ≲ U.
+Proof. move => > -> ->. hauto l:on use:renaming. Qed.
+
+Lemma weakening_su : forall n Γ (A0 A1 B : PTm n) i,
+    Γ ⊢ B ∈ PUniv i ->
+    Γ ⊢ A0 ≲ A1 ->
+    funcomp (ren_PTm shift) (scons B Γ) ⊢ ren_PTm shift A0 ≲ ren_PTm shift A1.
+Proof.
+  move => n Γ 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.
+Proof. hauto lq:on use:regularity. Qed.
+
+Lemma Su_Pi_Proj2_Var n Γ (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A1 B1 ->
+  funcomp (ren_PTm shift) (scons A1 Γ) ⊢ B0 ≲ B1.
+Proof.
+  move => h.
+  have /Su_Pi_Proj1 h1 := h.
+  have /regularity_sub0 [i h2] := h1.
+  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.
+  sfirstorder use:wff_mutual.
+  by asimpl.
+  by asimpl.
+Qed.

From ba77752329c2b516febcd86d8ebeb54e98a80cb5 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 12 Feb 2025 20:18:12 -0500
Subject: [PATCH 087/210] Add more cases to the soundness proof

---
 theories/algorithmic.v | 33 ++++++++++++++++++---------------
 theories/structural.v  | 16 ++++++++++++++++
 2 files changed, 34 insertions(+), 15 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 34b60f8..adde1e5 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -232,8 +232,8 @@ Proof.
     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 admit.
-    have ? : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
-    have ? : Γ ⊢ PBind PPi A1 B1 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
+    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 [l ?] : exists j, Γ ⊢ A2 ∈ PUniv j by hauto lq:on rew:off use:regularity, Bind_Inv.
@@ -242,21 +242,21 @@ Proof.
     eauto using E_Symmetric, Su_Eq.
     apply : E_Abs; eauto. hauto l:on use:regularity.
     apply iha.
+    move /Su_Pi_Proj2_Var in hp0.
+    apply : T_Conv; eauto.
     eapply ctx_eq_subst_one with (A0 := A0); eauto.
-    admit.
-    admit.
-    admit.
+    move /Su_Pi_Proj2_Var in hp1.
+    apply : T_Conv; eauto.
     eapply ctx_eq_subst_one with (A0 := A1); eauto.
-    admit.
-    admit.
-    admit.
-    (* Need to use the fundamental lemma to show that U normalizes to a Pi type *)
   - abstract : hAppL.
     move => n 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 admit.
     have hPi'' : Γ ⊢ PBind PPi A2 B2 ≲ A by eauto using Su_Eq, Su_Transitive, E_Symmetric.
+    have [j0 ?] : exists j0, Γ ⊢ A0 ∈ PUniv j0 by move /regularity_sub0 in hPi; hauto lq:on use:Bind_Inv.
+    have [j2 ?] : exists j0, Γ ⊢ A2 ∈ PUniv j0 by move /regularity_sub0 in hPi''; hauto lq:on use:Bind_Inv.
     have hPi' : Γ ⊢ PBind PPi A0 B0 ≲ PBind PPi A2 B2 by eauto using Su_Eq, Su_Transitive.
+    have hPidup := hPi'.
     apply E_Conv with (A := PBind PPi A2 B2); eauto.
     have /regularity_sub0 [i' hPi0] := hPi.
     have : Γ ⊢ PAbs (PApp (ren_PTm shift u) (VarPTm var_zero)) ≡ u ∈ PBind PPi A2 B2.
@@ -270,13 +270,16 @@ Proof.
       by eauto using E_Transitive.
     apply : E_Abs; eauto. hauto l:on use:regularity.
     apply iha.
-    admit.
+    move /Su_Pi_Proj2_Var in hPi'.
+    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 weakening_wt' with (a := u) (A := PBind PPi A2 B2). reflexivity. by asimpl.
-    admit.
-    apply : T_Conv; eauto. apply : Su_Eq; eauto.
-    apply T_Var. apply : Wff_Cons'; eauto.
-    admit.
+    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.
   (* Mirrors the last case *)
   - move => n a u hu ha iha Γ A hu0 ha0.
     apply E_Symmetric.
diff --git a/theories/structural.v b/theories/structural.v
index 16fc334..2773c3c 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -658,3 +658,19 @@ Proof.
   by asimpl.
   by asimpl.
 Qed.
+
+Lemma Su_Sig_Proj2_Var n Γ (A0 A1 : PTm n) B0 B1 :
+  Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A1 B1 ->
+  funcomp (ren_PTm shift) (scons A0 Γ) ⊢ B0 ≲ B1.
+Proof.
+  move => h.
+  have /Su_Sig_Proj1 h1 := h.
+  have /regularity_sub0 [i h2] := h1.
+  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.
+  sfirstorder use:wff_mutual.
+  by asimpl.
+  by asimpl.
+Qed.

From 1d3920fce16e169b195ed1fb8f2f9faad6d4086c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 12 Feb 2025 21:13:47 -0500
Subject: [PATCH 088/210] Prove the case for pair and neutral

---
 theories/algorithmic.v | 39 ++++++++++++++++++++++++++++++++++-----
 1 file changed, 34 insertions(+), 5 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index adde1e5..ceaa805 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -291,22 +291,51 @@ Proof.
     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 admit.
     apply E_Conv with (A := PBind PSig A2 B2); last by eauto using E_Symmetric, Su_Eq.
+    have h0 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 by eauto using Su_Transitive, Su_Eq, E_Symmetric.
+    have h1 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2 by eauto using Su_Transitive, Su_Eq, E_Symmetric.
+    have /Su_Sig_Proj1 h0' := h0.
+    have /Su_Sig_Proj1 h1' := h1.
+    move => [:eqa].
     apply : E_Pair; eauto. hauto l:on use:regularity.
-    apply iha. admit. admit.
-    admit.
+    abstract : eqa. apply iha; eauto using T_Conv.
+    apply ihb.
+    + apply T_Conv with (A := subst_PTm (scons a0 VarPTm) B0); eauto.
+      have : Γ ⊢ a0 ≡ a0 ∈A0 by eauto using E_Refl.
+      hauto l:on use:Su_Sig_Proj2.
+    + apply T_Conv with (A := subst_PTm (scons a1 VarPTm) B2); eauto; cycle 1.
+      move /E_Symmetric in eqa.
+      have ? : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i by hauto use:regularity.
+      apply:bind_inst; eauto.
+      apply : T_Conv; eauto.
+      have : Γ ⊢ a1 ≡ a1 ∈ A1 by eauto using E_Refl.
+      hauto l:on use:Su_Sig_Proj2.
   - move => {hAppL}.
     abstract : hPairL.
+    move => {hAppL}.
     move => n 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 admit.
     have hA' : Γ ⊢ PBind PSig A2 B2 ≲ A by eauto using E_Symmetric, Su_Eq.
     move /E_Conv : (hA'). apply.
+    have hSig : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 by eauto using E_Symmetric, Su_Eq, Su_Transitive.
+    have hA02 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Sig_Proj1.
+    have hu'  : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+    move => [:ih0'].
     apply : E_Transitive; last (apply E_Symmetric; apply : E_PairEta).
     apply : E_Pair; eauto. hauto l:on use:regularity.
-    apply ih0. admit. admit.
-    apply ih1. admit. admit.
+    abstract : ih0'.
+    apply ih0. apply : T_Conv; eauto.
+    by eauto using T_Proj1.
+    apply ih1. apply : T_Conv; eauto.
+    move /E_Refl in ha0.
+    hauto l:on use:Su_Sig_Proj2.
+    move /T_Proj2 in hu'.
+    apply : T_Conv; eauto.
+    move /E_Symmetric in ih0'.
+    move /regularity_sub0 in hA'.
+    hauto l:on use:bind_inst.
     hauto l:on use:regularity.
-    apply : T_Conv; eauto using Su_Eq.
+    eassumption.
   (* Same as before *)
   - move {hAppL}.
     move => *. apply E_Symmetric. apply : hPairL;

From 1f1b8a83db6d2fc76dae04a3b339b619a495f4ec Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 12 Feb 2025 22:00:44 -0500
Subject: [PATCH 089/210] Pull out some inversion lemmas to prove later

---
 theories/algorithmic.v | 49 +++++++++++++++++++++++++++++++-----------
 1 file changed, 36 insertions(+), 13 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index ceaa805..3046ced 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -158,6 +158,26 @@ Lemma coqeq_symmetric_mutual : forall n,
     (forall (a b : PTm n), a ⇔ b -> b ⇔ a).
 Proof. apply coqeq_mutual; qauto l:on ctrs:CoqEq,CoqEq_R, CoqEq_Neu. Qed.
 
+Lemma Sub_Bind_InvR n Γ p (A : PTm n) B C :
+  Γ ⊢ PBind p A B ≲ C ->
+  exists i A0 B0, Γ ⊢ C ≡ PBind p A0 B0 ∈ PUniv i.
+Proof.
+Admitted.
+
+Lemma Sub_Bind_InvL n Γ p (A : PTm n) B C :
+  Γ ⊢ C ≲ PBind p A B ->
+  exists i A0 B0, Γ ⊢ PBind p A0 B0 ≡ C ∈ PUniv i.
+Proof.
+Admitted.
+
+Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
+  Γ ⊢ PUniv i ≲ C ->
+  exists j, Γ ⊢ C ≡ PUniv j ∈ PUniv (S j).
+Proof.
+Admitted.
+
+(* Lemma Sub_Univ_InvR *)
+
 Lemma coqeq_sound_mutual : forall n,
     (forall (a b : PTm n), a ∼ b -> forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> exists C,
        Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ≡ b ∈ C) /\
@@ -180,7 +200,7 @@ Proof.
       specialize ihu with (1 := hu0') (2 := hu1').
       move : ihu.
       move => [C][ih0][ih1]ih.
-      have [i[A2[B2 h2]]] : exists i A2 B2, Γ ⊢ PBind PSig A2 B2 ≡ C ∈ PUniv i by admit.
+      have [i[A2[B2 h2]]] : exists i A2 B2, Γ ⊢ PBind PSig A2 B2 ≡ C ∈ PUniv i by sfirstorder use:Sub_Bind_InvL.
       exists  A2.
       have [h3 h4] : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 by qauto l:on use:Su_Eq, Su_Transitive.
       repeat split;
@@ -193,7 +213,7 @@ Proof.
       specialize ihu with (1 := hu0') (2 := hu1').
       move : ihu.
       move => [C][ih0][ih1]ih.
-      have [A2 [B2 [i hi]]] : exists A2 B2 i, Γ ⊢ PBind PSig A2 B2 ≡ C ∈ PUniv i by admit.
+      have [A2 [B2 [i hi]]] : exists A2 B2 i, Γ ⊢ PBind PSig A2 B2 ≡ C ∈ PUniv i by hauto q:on use:Sub_Bind_InvL.
       have [h3 h4] : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 /\ Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 by qauto l:on use:Su_Eq, Su_Transitive.
       have h5 : Γ ⊢ u0 ≡ u1 ∈ PBind PSig A2 B2 by eauto using E_Conv_E.
       exists (subst_PTm (scons (PProj PL u0) VarPTm) B2).
@@ -211,7 +231,7 @@ Proof.
     move /App_Inv : wta1 => [A1][B1][hu1][ha1]hU1.
     move : ihu hu0 hu1. repeat move/[apply].
     move => [C][hC0][hC1]hu01.
-    have [i [A2 [B2 hPi]]] : exists i A2 B2, Γ ⊢ PBind PPi A2 B2 ≡ C ∈ PUniv i by admit.
+    have [i [A2 [B2 hPi]]] : exists i A2 B2, Γ ⊢ PBind PPi A2 B2 ≡ C ∈ PUniv i by sfirstorder use:Sub_Bind_InvL.
     have ? : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by eauto using Su_Eq, Su_Transitive.
     have h : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by eauto using Su_Eq, Su_Transitive.
     have ha' : Γ ⊢ a0 ≡ a1 ∈ A2 by
@@ -231,7 +251,7 @@ Proof.
   - move => n 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 admit.
+    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.
@@ -251,7 +271,7 @@ Proof.
   - abstract : hAppL.
     move => n 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 admit.
+    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.
     have [j0 ?] : exists j0, Γ ⊢ A0 ∈ PUniv j0 by move /regularity_sub0 in hPi; hauto lq:on use:Bind_Inv.
     have [j2 ?] : exists j0, Γ ⊢ A2 ∈ PUniv j0 by move /regularity_sub0 in hPi''; hauto lq:on use:Bind_Inv.
@@ -289,7 +309,7 @@ Proof.
   - move => {hAppL hPairL} n 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 admit.
+    have [i[A2[B2 h2]]] : exists i A2 B2, Γ ⊢ A ≡ PBind PSig A2 B2 ∈ PUniv i by hauto l:on use:Sub_Bind_InvR.
     apply E_Conv with (A := PBind PSig A2 B2); last by eauto using E_Symmetric, Su_Eq.
     have h0 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 by eauto using Su_Transitive, Su_Eq, E_Symmetric.
     have h1 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2 by eauto using Su_Transitive, Su_Eq, E_Symmetric.
@@ -314,7 +334,7 @@ Proof.
     move => {hAppL}.
     move => n 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 admit.
+    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.
     move /E_Conv : (hA'). apply.
     have hSig : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 by eauto using E_Symmetric, Su_Eq, Su_Transitive.
@@ -344,20 +364,23 @@ Proof.
   - move => {hAppL hPairL} n 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 by admit.
+    have [l [k hk]] : exists l k, Γ ⊢ A ≡ PUniv k ∈ PUniv l
+        by hauto lq:on use:Sub_Univ_InvR.
+    have hjk : Γ ⊢ PUniv j ≲ PUniv k by eauto using Su_Eq, Su_Transitive.
+    have hik : Γ ⊢ PUniv i ≲ PUniv k by eauto using Su_Eq, Su_Transitive.
     apply E_Conv with (A := PUniv k); last by eauto using Su_Eq, E_Symmetric.
     move => [:eqA].
     apply E_Bind. abstract : eqA. apply ihA.
     apply T_Conv with (A := PUniv i); eauto.
-    by eauto using Su_Transitive, Su_Eq, E_Symmetric.
     apply T_Conv with (A := PUniv j); eauto.
-    by eauto using Su_Transitive, Su_Eq, E_Symmetric.
     apply ihB.
-    apply T_Conv with (A := PUniv i); eauto. admit.
+    apply T_Conv with (A := PUniv i); eauto.
+    move : weakening_su hik hA0. by repeat move/[apply].
     apply T_Conv with (A := PUniv j); eauto.
-    apply : ctx_eq_subst_one; eauto. apply : Su_Eq; apply eqA.  admit.
+    apply : ctx_eq_subst_one; eauto. apply : Su_Eq; apply eqA.
+    move : weakening_su hjk hA0. by repeat move/[apply].
   - hauto lq:on ctrs:Eq,LEq,Wt.
   - move => n 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.
-Admitted.
+Qed.

From 0a70be3e57ec7dda63575eba72ed79c8221a617a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 13 Feb 2025 17:09:58 -0500
Subject: [PATCH 090/210] Define the size metric for the completeness proof

---
 theories/algorithmic.v | 5 ++++-
 1 file changed, 4 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 3046ced..819878e 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -3,7 +3,7 @@ Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
 From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
 Require Import Psatz.
-From stdpp Require Import relations (rtc(..)).
+From stdpp Require Import relations (rtc(..), nsteps(..)).
 Require Import Coq.Logic.FunctionalExtensionality.
 
 Module HRed.
@@ -384,3 +384,6 @@ Proof.
     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 (a b : PTm n) :=
+  exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j a vb /\ size_PTm va + size_PTm vb + i + j = n.

From 849169be7fc3c9930841b9f42cb6865a2e552143 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 13 Feb 2025 17:46:35 -0500
Subject: [PATCH 091/210] Start coqeq_complete

---
 theories/algorithmic.v | 16 +++++++++++++++-
 1 file changed, 15 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 819878e..03d5eaa 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -386,4 +386,18 @@ Proof.
 Qed.
 
 Definition algo_metric n (a b : PTm n) :=
-  exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j a vb /\ size_PTm va + size_PTm vb + i + j = n.
+  exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j a vb /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j = n.
+
+Lemma ishne_red n (a : PTm n) : SN a -> ~ ishne a -> exists b, HRed.R a b.
+Admitted.
+
+Lemma coqeq_complete n (a b : PTm n) :
+  algo_metric n a b -> DJoin.R 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).
+Proof.
+  move : n a b.
+  elim /Wf_nat.lt_wf_ind.
+  move => n ih.
+  move => a.
+  case /orP : (orNb (ishf a)).

From 093fc8f9cba936a741eb5a7ba6a8f7ae962b470f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 14 Feb 2025 13:29:44 -0500
Subject: [PATCH 092/210] Finish algo_metric_case

---
 theories/algorithmic.v | 43 +++++++++++++++++++++++++++++++++++++-----
 theories/common.v      |  1 +
 2 files changed, 39 insertions(+), 5 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 03d5eaa..9805cb6 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -5,6 +5,7 @@ Require Import ssreflect ssrbool.
 Require Import Psatz.
 From stdpp Require Import relations (rtc(..), nsteps(..)).
 Require Import Coq.Logic.FunctionalExtensionality.
+Require Import Cdcl.Itauto.
 
 Module HRed.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
@@ -385,11 +386,43 @@ Proof.
     hauto lq:on use:HReds.ToEq, E_Symmetric, E_Transitive.
 Qed.
 
-Definition algo_metric n (a b : PTm n) :=
-  exists i j va vb, nsteps LoRed.R i a va /\ nsteps LoRed.R j a vb /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j = n.
+Definition algo_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.
 
-Lemma ishne_red n (a : PTm n) : SN a -> ~ ishne a -> exists b, HRed.R a b.
-Admitted.
+Lemma ne_hne n (a : PTm n) : ne a -> ishne a.
+Proof. elim : a => //=; sfirstorder b:on. Qed.
+
+Lemma hf_hred_lored n (a b : PTm n) :
+  ~~ ishf a ->
+  LoRed.R a b ->
+  HRed.R a b \/ ishne a.
+Proof.
+  move => + h. elim : n a b/ h=>//=.
+  - hauto l:on use:HRed.AppAbs.
+  - hauto l:on use:HRed.ProjPair.
+  - hauto lq:on ctrs:HRed.R.
+  - sfirstorder use:ne_hne.
+  - hauto lq:on ctrs:HRed.R.
+Qed.
+
+Lemma algo_metric_case n k (a b : PTm n) :
+  algo_metric k a b ->
+  ishf a \/ ishne a \/ exists k' a', HRed.R a a' /\ algo_metric k' a' b /\ k' < k.
+Proof.
+  move=>[i][j][va][vb][h0][h1][h2][h3]h4.
+  case : a h0 => //=; try firstorder.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.AppAbs unfold:algo_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder qb:on use:ne_hne.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+Qed.
 
 Lemma coqeq_complete n (a b : PTm n) :
   algo_metric n a b -> DJoin.R a b ->
@@ -400,4 +433,4 @@ Proof.
   elim /Wf_nat.lt_wf_ind.
   move => n ih.
   move => a.
-  case /orP : (orNb (ishf a)).
+  case /orP : (orNb (ishf a || ishne a)).
diff --git a/theories/common.v b/theories/common.v
index 6ad3d44..ac70322 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -60,6 +60,7 @@ Fixpoint ishne {n} (a : PTm n) :=
   | VarPTm _ => true
   | PApp a _ => ishne a
   | PProj _ a => ishne a
+  | PBot => true
   | _ => false
   end.
 

From 5ed366f093a269bf81af2116786977af6b940e1e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 14 Feb 2025 14:49:19 -0500
Subject: [PATCH 093/210] Prove some easy cases of completeness

---
 theories/algorithmic.v | 75 +++++++++++++++++++++++++++++++++++++-----
 1 file changed, 66 insertions(+), 9 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 9805cb6..97a7a3a 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -146,6 +146,14 @@ 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)  :
+  HRed.R a a' ->
+  a' ⇔ b ->
+  a ⇔ b.
+Proof.
+  hauto lq:on ctrs:rtc, CoqEq_R inv:CoqEq_R.
+Qed.
+
 Scheme
   coqeq_neu_ind := Induction for CoqEq_Neu Sort Prop
   with coqeq_ind := Induction for CoqEq Sort Prop
@@ -387,7 +395,7 @@ Proof.
 Qed.
 
 Definition algo_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.
+  exists i j va vb v, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ rtc ERed.R va v /\ rtc ERed.R vb v /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j = k.
 
 Lemma ne_hne n (a : PTm n) : ne a -> ishne a.
 Proof. elim : a => //=; sfirstorder b:on. Qed.
@@ -404,12 +412,11 @@ Proof.
   - sfirstorder use:ne_hne.
   - hauto lq:on ctrs:HRed.R.
 Qed.
-
 Lemma algo_metric_case n k (a b : PTm n) :
   algo_metric k a b ->
-  ishf a \/ ishne a \/ exists k' a', HRed.R a a' /\ algo_metric k' a' b /\ k' < k.
+  (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ algo_metric k' a' b /\ k' < k.
 Proof.
-  move=>[i][j][va][vb][h0][h1][h2][h3]h4.
+  move=>[i][j][va][vb][v][h0][h1][h2][h3][h4][h5]h6.
   case : a h0 => //=; try firstorder.
   - inversion h0 as [|A B C D E F]; subst.
     hauto qb:on use:ne_hne.
@@ -424,13 +431,63 @@ Proof.
     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
 Qed.
 
-Lemma coqeq_complete n (a b : PTm n) :
-  algo_metric n a b -> DJoin.R a b ->
+Lemma algo_metric_sym n k (a b : PTm n) :
+  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) :
+  HRed.R a b ->
+  ishne a ->
+  False.
+Proof. induction 1; sfirstorder. Qed.
+
+Lemma hf_not_hne n (a : PTm n) :
+  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 :
+  Γ ⊢ PAbs a ∈ A ->
+  Γ ⊢ PPair b0 b1 ∈ A ->
+  False.
+Proof.
+  move /Abs_Inv => [A0][B0][_]haU.
+  move /Pair_Inv => [A1][B1][_][_]hbU.
+  move /Sub_Bind_InvR : haU => [i][A2][B2]h2.
+  have : Γ ⊢ PBind PSig A1 B1 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
+  clear.
+Admitted.
+
+Lemma coqeq_complete n k (a b : PTm n) :
+  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).
 Proof.
-  move : n a b.
+  move : k n a b.
   elim /Wf_nat.lt_wf_ind.
   move => n ih.
-  move => a.
-  case /orP : (orNb (ishf a || ishne a)).
+  move => k 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.
+  move : k 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.
+  (* Cases where a and b can't take wh steps *)
+  move {hstepL}.
+  move : k a b h. move => [:hAppNeu hProjNeu].
+  move => k a b h.
+  case => fb; case => fa.
+  - split; last by sfirstorder use:hf_not_hne.
+    case : a h fb fa => //=.
+    + case : b => //=.

From ea14ba96024425172df7921251b682d5a090d5a1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 14 Feb 2025 16:17:02 -0500
Subject: [PATCH 094/210] Prove most of cases of AbsAbs

---
 theories/algorithmic.v | 134 ++++++++++++++++++++++++++++++++++-------
 theories/soundness.v   |   3 +
 2 files changed, 114 insertions(+), 23 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 97a7a3a..d2ef828 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1,5 +1,5 @@
 Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
-  common typing preservation admissible fp_red structural.
+  common typing preservation admissible fp_red structural soundness.
 From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
 Require Import Psatz.
@@ -59,6 +59,49 @@ Lemma E_Conv_E n Γ (a b : PTm n) A B i :
   Γ ⊢ a ≡ b ∈ B.
 Proof. qauto use:E_Conv, Su_Eq, E_Symmetric. Qed.
 
+Lemma Sub_Bind_InvR n Γ p (A : PTm n) B C :
+  Γ ⊢ PBind p A B ≲ C ->
+  exists i A0 B0, Γ ⊢ C ≡ PBind p A0 B0 ∈ PUniv i.
+Proof.
+Admitted.
+
+Lemma Sub_Bind_InvL n Γ p (A : PTm n) B C :
+  Γ ⊢ C ≲ PBind p A B ->
+  exists i A0 B0, Γ ⊢ PBind p A0 B0 ≡ C ∈ PUniv i.
+Proof.
+Admitted.
+
+Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
+  Γ ⊢ PUniv i ≲ C ->
+  exists j, Γ ⊢ C ≡ PUniv j ∈ PUniv (S j).
+Proof.
+Admitted.
+
+Lemma T_Abs_Inv n Γ (a0 a1 : PTm (S n)) U :
+  Γ ⊢ PAbs a0 ∈ U ->
+  Γ ⊢ PAbs a1 ∈ U ->
+  exists Δ V, Δ ⊢ a0 ∈ V /\ Δ ⊢ a1 ∈ V.
+Proof.
+  move /Abs_Inv => [A0][B0][wt0]hU0.
+  move /Abs_Inv => [A1][B1][wt1]hU1.
+  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.
+  have [k ?] : exists k, Γ ⊢ A2 ∈ PUniv k by hauto lq:on use:regularity, Bind_Inv.
+  split.
+  - have /Su_Pi_Proj2_Var ? := hSu'.
+    have /Su_Pi_Proj1 ? := hSu'.
+    move /regularity_sub0 : hSu' => [j] /Bind_Inv [k0 [? ?]].
+    apply T_Conv with (A := B0); eauto.
+    apply : ctx_eq_subst_one; eauto.
+  - have /Su_Pi_Proj2_Var ? := hSu.
+    have /Su_Pi_Proj1 ? := hSu.
+    move /regularity_sub0 : hSu => [j] /Bind_Inv [k0 [? ?]].
+    apply T_Conv with (A := B1); eauto.
+    apply : ctx_eq_subst_one; eauto.
+Qed.
+
 (* Coquand's algorithm with subtyping *)
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
@@ -167,24 +210,6 @@ Lemma coqeq_symmetric_mutual : forall n,
     (forall (a b : PTm n), a ⇔ b -> b ⇔ a).
 Proof. apply coqeq_mutual; qauto l:on ctrs:CoqEq,CoqEq_R, CoqEq_Neu. Qed.
 
-Lemma Sub_Bind_InvR n Γ p (A : PTm n) B C :
-  Γ ⊢ PBind p A B ≲ C ->
-  exists i A0 B0, Γ ⊢ C ≡ PBind p A0 B0 ∈ PUniv i.
-Proof.
-Admitted.
-
-Lemma Sub_Bind_InvL n Γ p (A : PTm n) B C :
-  Γ ⊢ C ≲ PBind p A B ->
-  exists i A0 B0, Γ ⊢ PBind p A0 B0 ≡ C ∈ PUniv i.
-Proof.
-Admitted.
-
-Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
-  Γ ⊢ PUniv i ≲ C ->
-  exists j, Γ ⊢ C ≡ PUniv j ∈ PUniv (S j).
-Proof.
-Admitted.
-
 (* Lemma Sub_Univ_InvR *)
 
 Lemma coqeq_sound_mutual : forall n,
@@ -395,7 +420,7 @@ Proof.
 Qed.
 
 Definition algo_metric {n} k (a b : PTm n) :=
-  exists i j va vb v, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ rtc ERed.R va v /\ rtc ERed.R vb v /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j = k.
+  exists i j va vb v, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ rtc ERed.R va v /\ rtc ERed.R vb v /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j <= k.
 
 Lemma ne_hne n (a : PTm n) : ne a -> ishne a.
 Proof. elim : a => //=; sfirstorder b:on. Qed.
@@ -461,8 +486,54 @@ Proof.
   move /Pair_Inv => [A1][B1][_][_]hbU.
   move /Sub_Bind_InvR : haU => [i][A2][B2]h2.
   have : Γ ⊢ PBind PSig A1 B1 ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
-  clear.
-Admitted.
+  clear. move /synsub_to_usub. hauto l:on use:Sub.bind_inj.
+Qed.
+
+Lemma Univ_Inv n Γ i U :
+  Γ ⊢ @PUniv n 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 //=.
+  - hauto l:on use:E_Refl, Su_Eq, T_Univ.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+Lemma T_AbsUniv_Imp n Γ a i (A : PTm n) :
+  Γ ⊢ PAbs a ∈ A ->
+  Γ ⊢ PUniv i ∈ A ->
+  False.
+Proof.
+  move /Abs_Inv => [A0][B0][_]haU.
+  move /Univ_Inv => [h0 h1].
+  move /Sub_Univ_InvR : h1 => [j hU].
+  have : Γ ⊢ PBind PPi A0 B0 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
+  clear. move /synsub_to_usub.
+  hauto lq:on use:Sub.bind_univ_noconf.
+Qed.
+
+Lemma T_AbsBind_Imp n Γ a p A0 B0 (U : PTm n) :
+  Γ ⊢ PAbs a ∈ U ->
+  Γ ⊢ PBind p A0 B0 ∈ U ->
+  False.
+Proof.
+  move /Abs_Inv => [A1][B1][_ ha].
+  move /Bind_Inv => [i [_ [_ h]]].
+  move /Sub_Univ_InvR : h => [j hU].
+  have : Γ ⊢ PBind PPi A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
+  clear. move /synsub_to_usub.
+  hauto lq:on use:Sub.bind_univ_noconf.
+Qed.
+
+Lemma lored_nsteps_abs_inv k n (a : PTm (S n)) 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.
+  move : a E.
+  elim : u b /hu.
+  - hauto l:on.
+  - hauto lq:on ctrs:nsteps inv:LoRed.R.
+Qed.
 
 Lemma coqeq_complete n k (a b : PTm n) :
   algo_metric k a b ->
@@ -485,9 +556,26 @@ Proof.
   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 => [:hAppNeu hProjNeu].
+  move : k a b h. (* move => [:hAppNeu hProjNeu]. *)
   move => k a b h.
   case => fb; case => fa.
   - split; last by sfirstorder use:hf_not_hne.
     case : a h fb fa => //=.
+    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp.
+      move => a0 a1 ha0 _ _ Γ A wt0 wt1.
+      move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]].
+      apply : CE_HRed; eauto using rtc_refl.
+      apply CE_AbsAbs.
+      suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto.
+      have ? : n > 0 by sauto solve+:lia.
+      exists (n - 1). split; first by lia.
+      move : ha0. rewrite /algo_metric.
+      move => [i][j][va][vb][v][hr0][hr1][hr0'][hr1'][nfva][nfvb]har.
+      apply lored_nsteps_abs_inv in hr0, hr1.
+      move : hr0 => [va' [hr00 hr01]].
+      move : hr1 => [vb' [hr10 hr11]]. move {ih}.
+      exists i,j,va',vb'. subst.
+      suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
+      exists v0. repeat split =>//=. simpl in har. lia.
+      admit.
     + case : b => //=.
diff --git a/theories/soundness.v b/theories/soundness.v
index 8dd87da..b11dded 100644
--- a/theories/soundness.v
+++ b/theories/soundness.v
@@ -10,3 +10,6 @@ Theorem fundamental_theorem :
   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.
+Proof. hauto lq:on rew:off use:fundamental_theorem, SemLEq_SN_Sub. Qed.

From 8fd691953845314e52d05c4fa177b0e79fe8ba74 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 14 Feb 2025 16:57:53 -0500
Subject: [PATCH 095/210] Make progress on coqeq_complete

---
 theories/algorithmic.v | 55 +++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 54 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index d2ef828..70ea3ab 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -489,6 +489,18 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_inj.
 Qed.
 
+Lemma T_PairBind_Imp n Γ (a0 a1 : PTm n) p A0 B0 U :
+  Γ ⊢ PPair a0 a1 ∈ U ->
+  Γ ⊢ PBind p A0 B0  ∈ U ->
+  False.
+Proof.
+  move /Pair_Inv => [A1][B1][_][_]hbU.
+  move /Bind_Inv => [i][_ [_ haU]].
+  move /Sub_Univ_InvR : haU => [j]hU.
+  have : Γ ⊢ PBind PSig A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
+  clear. move /synsub_to_usub. hauto l:on use:Sub.bind_univ_noconf.
+Qed.
+
 Lemma Univ_Inv n Γ i U :
   Γ ⊢ @PUniv n i ∈ U ->
   Γ ⊢ PUniv i ∈ PUniv (S i)  /\ Γ ⊢ PUniv (S i) ≲ U.
@@ -499,6 +511,19 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
+Lemma T_PairUniv_Imp n Γ (a0 a1 : PTm n) i U :
+  Γ ⊢ PPair a0 a1 ∈ U ->
+  Γ ⊢ PUniv i ∈ U ->
+  False.
+Proof.
+  move /Pair_Inv => [A1][B1][_][_]hbU.
+  move /Univ_Inv => [h0 h1].
+  move /Sub_Univ_InvR : h1 => [j hU].
+  have : Γ ⊢ PBind PSig A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
+  clear. move /synsub_to_usub.
+  hauto lq:on use:Sub.bind_univ_noconf.
+Qed.
+
 Lemma T_AbsUniv_Imp n Γ a i (A : PTm n) :
   Γ ⊢ PAbs a ∈ A ->
   Γ ⊢ PUniv i ∈ A ->
@@ -556,10 +581,11 @@ Proof.
   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 => [:hAppNeu hProjNeu]. *)
+  move : k a b h. move => [:hnfneu].
   move => k a b h.
   case => fb; case => fa.
   - split; last by sfirstorder use:hf_not_hne.
+    move {hnfneu}.
     case : a h fb fa => //=.
     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp.
       move => a0 a1 ha0 _ _ Γ A wt0 wt1.
@@ -578,4 +604,31 @@ Proof.
       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
       exists v0. repeat split =>//=. simpl in har. lia.
       admit.
+    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
+      move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
+      apply : CE_HRed; eauto using rtc_refl.
+      admit.
+      (* suff : exists l, l < n /\ algo_metric l a0 b0 /\ algo_metric l a1 b1. *)
+      (* move => [l [hl [hal0 hal1]]]. *)
+      (* apply CE_PairPair. eapply ih; eauto. *)
+      (* by eapply ih; eauto. *)
+    + admit.
+    + admit.
+  - move : k a b h fb fa. abstract : hnfneu.
+    admit.
+  - 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)).
+    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 => i j hi _ _.
+      * rewrite /algo_metric in hi.
+Admitted.

From 186f2138e60347e91fd8edede63a0f67a0f871dd Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 14 Feb 2025 19:08:41 -0500
Subject: [PATCH 096/210] Finish the var base case

---
 theories/algorithmic.v | 44 +++++++++++++++++++++++++++++++++++++++++-
 theories/fp_red.v      | 12 ++++++++++++
 2 files changed, 55 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 70ea3ab..d619282 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -550,6 +550,13 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
+Lemma T_Bot_Imp n Γ (A : PTm n) :
+  Γ ⊢ PBot ∈ A -> False.
+  move E : PBot => u hu.
+  move : E.
+  induction hu => //=.
+Qed.
+
 Lemma lored_nsteps_abs_inv k n (a : PTm (S n)) b :
   nsteps LoRed.R k (PAbs a) b -> exists b', nsteps LoRed.R k a b' /\ b = PAbs b'.
 Proof.
@@ -560,6 +567,19 @@ Proof.
   - hauto lq:on ctrs:nsteps inv:LoRed.R.
 Qed.
 
+Lemma algo_metric_join n k (a b : PTm n) :
+  algo_metric k a b ->
+  DJoin.R a b.
+  rewrite /algo_metric.
+  move => [i][j][va][vb][v][h0][h1][h2][h3]h4.
+  have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
+  have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
+  apply REReds.FromEReds in h2,h3.
+  apply LoReds.ToRReds in h0,h1.
+  apply REReds.FromRReds in h0,h1.
+  rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
+Qed.
+
 Lemma coqeq_complete n k (a b : PTm n) :
   algo_metric k a b ->
   (forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) /\
@@ -630,5 +650,27 @@ Proof.
     case : a h fb fa => //=.
     + case : b => //=.
       move => i j hi _ _.
-      * rewrite /algo_metric in hi.
+      * have ? : j = i by hauto lq:on use:algo_metric_join, DJoin.var_inj. subst.
+        move => Γ A B hA hB.
+        split. apply CE_VarCong.
+        exists (Γ i). hauto l:on use:Var_Inv.
+      * admit.
+      * admit.
+      * sfirstorder use:T_Bot_Imp.
+    + case : b => //=.
+      * admit.
+      (* real case *)
+      * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
+        move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
+        move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1.
+        admit.
+      * admit.
+      * sfirstorder use:T_Bot_Imp.
+    + case : b => //=.
+      * admit.
+      * admit.
+      (* real case *)
+      * admit.
+      * sfirstorder use:T_Bot_Imp.
+    + sfirstorder use:T_Bot_Imp.
 Admitted.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index beed0bb..3211325 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1686,6 +1686,14 @@ Module REReds.
     hauto lq:on rew:off ctrs:rtc inv:RERed.R.
   Qed.
 
+  Lemma var_inv n (i : fin n) C :
+    rtc RERed.R (VarPTm i) C ->
+    C = VarPTm i.
+  Proof.
+    move E : (VarPTm i) => u hu.
+    move : i E. elim : u C /hu; hauto lq:on rew:off 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).
   Proof.
@@ -2188,6 +2196,10 @@ Module DJoin.
     hauto lq:on rew:off use:REReds.bind_inv.
   Qed.
 
+  Lemma var_inj n (i j : fin n) :
+    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.
   Proof.

From 8d765c495d12b7ff074355b5ad22c751c659b023 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 14 Feb 2025 20:41:56 -0500
Subject: [PATCH 097/210] Add some more injection lemmas for neutrals

---
 theories/algorithmic.v | 53 ++++++++++++++++++++++++++
 theories/fp_red.v      | 85 +++++++++++++++++++++++++++++++++++++++++-
 2 files changed, 137 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index d619282..84472bc 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -567,6 +567,53 @@ Proof.
   - hauto lq:on ctrs:nsteps inv:LoRed.R.
 Qed.
 
+Lemma lored_hne_preservation n (a b : PTm n) :
+    LoRed.R a b -> ishne a -> ishne b.
+Proof.  induction 1; sfirstorder. Qed.
+
+Lemma lored_nsteps_app_inv k n (a0 b0 C : PTm n) :
+    nsteps LoRed.R k (PApp a0 b0) C ->
+    ishne a0 ->
+    exists i j a1 b1,
+      i <= k /\ j <= k /\
+        C = PApp a1 b1 /\
+        nsteps LoRed.R i a0 a1 /\
+        nsteps LoRed.R j b0 b1.
+Proof.
+  move E : (PApp a0 b0) => u hu. move : a0 b0 E.
+  elim : k u C / hu.
+  - sauto lq:on.
+  - move => k a0 a1 a2 ha01 ha12 ih a3 b0 ?.  subst.
+    inversion ha01; subst => //=.
+    spec_refl.
+    move => h.
+    have : ishne a4 by sfirstorder use:lored_hne_preservation.
+    move : ih => /[apply]. move => [i][j][a1][b1][?][?][?][h0]h1.
+    subst. exists (S i),j,a1,b1. sauto lq:on solve+:lia.
+    spec_refl. move : ih => /[apply].
+    move => [i][j][a1][b1][?][?][?][h0]h1. subst.
+    exists i, (S j), a1, b1. sauto lq:on solve+:lia.
+Qed.
+
+Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
+  algo_metric k (PApp a0 b0) (PApp a1 b1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ algo_metric j a0 a1 /\ algo_metric j b0 b1.
+Proof.
+  move => [i][j][va][vb][v][h0][h1][h2][h3][h4][h5]h6.
+  move => hne0 hne1.
+  move : lored_nsteps_app_inv h0 (hne0);repeat move/[apply].
+  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
+  move : lored_nsteps_app_inv h1 (hne1);repeat move/[apply].
+  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+  simpl in *. exists (k - 1).
+  split. lia.
+  split.
+  + rewrite /algo_metric.
+    have : exists a4 b4, rtc ERed.R a2 a4 /\ ERed.R
+    exists i0,i1,a2,b2.
+
 Lemma algo_metric_join n k (a b : PTm n) :
   algo_metric k a b ->
   DJoin.R a b.
@@ -663,6 +710,12 @@ Proof.
       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
         move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1.
+        have : DJoin.R (PApp b0 a0) (PApp b1 a1)
+          by hauto l:on use:algo_metric_join.
+        move : DJoin.hne_app_inj (hne0) (hne1). repeat move/[apply].
+        move => [hjb hja].
+        split. apply CE_AppCong => //=.
+        eapply ih; eauto.
         admit.
       * admit.
       * sfirstorder use:T_Bot_Imp.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3211325..50d8491 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -20,7 +20,7 @@ Ltac2 spec_refl () :=
                try (specialize $h with (1 := eq_refl))
            end)  (Control.hyps ()).
 
-Ltac spec_refl := ltac2:(spec_refl ()).
+Ltac spec_refl := ltac2:(Control.enter spec_refl).
 
 Module EPar.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
@@ -1510,10 +1510,45 @@ Module EReds.
     induction 1; hauto lq:on ctrs:rtc use:ERed.substing.
   Qed.
 
+  Lemma app_inv n (a0 b0 C : PTm n) :
+    rtc ERed.R (PApp a0 b0) C ->
+    exists a1 b1, C = PApp a1 b1 /\
+               rtc ERed.R a0 a1 /\
+               rtc ERed.R b0 b1.
+  Proof.
+    move E : (PApp a0 b0) => u hu.
+    move : a0 b0 E.
+    elim : u C / hu.
+    - hauto lq:on ctrs:rtc.
+    - move => a b c ha ha' iha a0 b0 ?. subst.
+      hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R.
+  Qed.
+
+  Lemma proj_inv n p (a C : PTm n) :
+    rtc ERed.R (PProj p a) C ->
+    exists c, C = PProj p c /\ rtc ERed.R a c.
+  Proof.
+    move E : (PProj p a) => u hu.
+    move : p a E.
+    elim : u C /hu;
+      hauto q:on ctrs:rtc,ERed.R inv:ERed.R.
+  Qed.
+
 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.
+
+  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm n) :
+    R (PApp a0 b0) (PApp a1 b1) ->
+    R a0 a1 /\ R b0 b1.
+  Proof.
+    hauto lq:on use:EReds.app_inv.
+  Qed.
+
+End EJoin.
 
 Module RERed.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
@@ -1602,6 +1637,10 @@ Module RERed.
     hauto q:on use:ToBetaEta, FromBeta, FromEta, RRed.substing, ERed.substing.
   Qed.
 
+  Lemma hne_preservation n (a b : PTm n) :
+    RERed.R a b -> ishne a -> ishne b.
+  Proof.  induction 1; sfirstorder. Qed.
+
 End RERed.
 
 Module REReds.
@@ -1694,6 +1733,32 @@ 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) :
+    rtc RERed.R (PApp a0 b0) C ->
+    ishne a0 ->
+    exists a1 b1, C = PApp a1 b1 /\
+               rtc RERed.R a0 a1 /\
+               rtc RERed.R b0 b1.
+  Proof.
+    move E : (PApp a0 b0) => u hu.
+    move : a0 b0 E.
+    elim : u C / hu.
+    - hauto lq:on ctrs:rtc.
+    - 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.
+
+  Lemma hne_proj_inv n p (a C : PTm n) :
+    rtc RERed.R (PProj p a) C ->
+    ishne a ->
+    exists c, C = PProj p c /\ rtc RERed.R a c.
+  Proof.
+    move E : (PProj p a) => u hu.
+    move : p a E.
+    elim : u C /hu;
+      hauto q:on 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).
   Proof.
@@ -2206,6 +2271,24 @@ Module DJoin.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
+  Lemma hne_app_inj n (a0 b0 a1 b1 : PTm n) :
+    R (PApp a0 b0) (PApp a1 b1) ->
+    ishne a0 ->
+    ishne a1 ->
+    R a0 a1 /\ R b0 b1.
+  Proof.
+    hauto lq:on use:REReds.hne_app_inv.
+  Qed.
+
+  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm n) :
+    R (PProj p0 a0) (PProj p1 a1) ->
+    ishne a0 ->
+    ishne a1 ->
+    p0 = p1 /\ R a0 a1.
+  Proof.
+    sauto lq:on use:REReds.hne_proj_inv.
+  Qed.
+
   Lemma FromRRed0 n (a b : PTm n) :
     RRed.R a b -> R a b.
   Proof.

From f0c18fd77e9bcca504fd1290877b235b6f67330c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 14 Feb 2025 21:11:27 -0500
Subject: [PATCH 098/210] Finish the neutral app case

---
 theories/algorithmic.v | 68 ++++++++++++++++++++++++++++++++----------
 1 file changed, 52 insertions(+), 16 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 84472bc..f754383 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -420,7 +420,7 @@ Proof.
 Qed.
 
 Definition algo_metric {n} k (a b : PTm n) :=
-  exists i j va vb v, nsteps LoRed.R i a va /\ nsteps LoRed.R j b vb /\ rtc ERed.R va v /\ rtc ERed.R vb v /\ nf va /\ nf vb /\ size_PTm va + size_PTm vb + i + j <= k.
+  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.
 Proof. elim : a => //=; sfirstorder b:on. Qed.
@@ -441,7 +441,7 @@ Lemma algo_metric_case n k (a b : PTm n) :
   algo_metric k a b ->
   (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ algo_metric k' a' b /\ k' < k.
 Proof.
-  move=>[i][j][va][vb][v][h0][h1][h2][h3][h4][h5]h6.
+  move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6.
   case : a h0 => //=; try firstorder.
   - inversion h0 as [|A B C D E F]; subst.
     hauto qb:on use:ne_hne.
@@ -601,27 +601,37 @@ Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
   ishne a1 ->
   exists j, j < k /\ algo_metric j a0 a1 /\ algo_metric j b0 b1.
 Proof.
-  move => [i][j][va][vb][v][h0][h1][h2][h3][h4][h5]h6.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
   move => hne0 hne1.
   move : lored_nsteps_app_inv h0 (hne0);repeat move/[apply].
   move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
   move : lored_nsteps_app_inv h1 (hne1);repeat move/[apply].
   move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
   simpl in *. exists (k - 1).
+  move /andP : h2 => [h20 h21].
+  move /andP : h3 => [h30 h31].
+  move /EJoin.hne_app_inj : h4 => [h40 h41].
   split. lia.
   split.
   + rewrite /algo_metric.
-    have : exists a4 b4, rtc ERed.R a2 a4 /\ ERed.R
-    exists i0,i1,a2,b2.
+    exists i0,j0,a2,a3.
+    repeat split => //=.
+    sfirstorder b:on use:ne_nf.
+    sfirstorder b:on use:ne_nf.
+    lia.
+  + rewrite /algo_metric.
+    exists i1,j1,b2,b3.
+    repeat split => //=; sfirstorder b:on use:ne_nf.
+Qed.
 
 Lemma algo_metric_join n k (a b : PTm n) :
   algo_metric k a b ->
   DJoin.R a b.
   rewrite /algo_metric.
-  move => [i][j][va][vb][v][h0][h1][h2][h3]h4.
+  move => [i][j][va][vb][h0][h1][h2][h3][[v [h40 h41]]]h5.
   have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
   have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
-  apply REReds.FromEReds in h2,h3.
+  apply REReds.FromEReds in h40,h41.
   apply LoReds.ToRReds in h0,h1.
   apply REReds.FromRReds in h0,h1.
   rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
@@ -663,13 +673,14 @@ Proof.
       have ? : n > 0 by sauto solve+:lia.
       exists (n - 1). split; first by lia.
       move : ha0. rewrite /algo_metric.
-      move => [i][j][va][vb][v][hr0][hr1][hr0'][hr1'][nfva][nfvb]har.
+      move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har.
       apply lored_nsteps_abs_inv in hr0, hr1.
       move : hr0 => [va' [hr00 hr01]].
       move : hr1 => [vb' [hr10 hr11]]. move {ih}.
       exists i,j,va',vb'. subst.
       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
-      exists v0. repeat split =>//=. simpl in har. lia.
+      repeat split =>//=. sfirstorder.
+      simpl in *; by lia.
       admit.
     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
@@ -710,13 +721,38 @@ Proof.
       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
         move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1.
-        have : DJoin.R (PApp b0 a0) (PApp b1 a1)
-          by hauto l:on use:algo_metric_join.
-        move : DJoin.hne_app_inj (hne0) (hne1). repeat move/[apply].
-        move => [hjb hja].
-        split. apply CE_AppCong => //=.
-        eapply ih; eauto.
-        admit.
+        move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply].
+        move => [j [hj [hal0 hal1]]].
+        have /ih := hal0.
+        move /(_ hj).
+        move => [_ ihb].
+        move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply].
+        move => [hb01][C][hT0]hT1.
+        move /Sub_Bind_InvL : (hT0).
+        move => [i][A2][B2]hE.
+        have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by
+          eauto using Su_Eq, Su_Transitive.
+        have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by
+          eauto using Su_Eq, Su_Transitive.
+        have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
+        have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
+        have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
+        have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
+        move : ih (hj) hal1. repeat move/[apply].
+        move => [ih _].
+        move : ih (ha0') (ha1'); repeat move/[apply].
+        move => iha.
+        split. sauto lq:on.
+        exists (subst_PTm (scons a0 VarPTm) B2).
+        split.
+        apply : Su_Transitive; eauto.
+        move /E_Refl in ha0.
+        hauto l:on use:Su_Pi_Proj2.
+        apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
+        move /regularity_sub0 : hSu10 => [i0].
+        have : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
+        hauto l:on use:bind_inst.
+        hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
       * admit.
       * sfirstorder use:T_Bot_Imp.
     + case : b => //=.

From 300530a93fc806ab3c4dfbe315cff16d56567e9f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 14 Feb 2025 21:31:27 -0500
Subject: [PATCH 099/210] Finish off some easy contradictory cases

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

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index f754383..5aa0a3b 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -712,11 +712,14 @@ Proof.
         move => Γ A B hA hB.
         split. apply CE_VarCong.
         exists (Γ i). hauto l:on use:Var_Inv.
-      * admit.
-      * admit.
+      * move => p p0 f /algo_metric_join. clear => ? ? _. exfalso.
+        hauto l:on use:REReds.var_inv, REReds.hne_app_inv.
+      * move => a0 a1 i /algo_metric_join. clear => ? ? _. exfalso.
+        hauto l:on use:REReds.var_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
     + case : b => //=.
-      * admit.
+      * clear. move => i a a0 /algo_metric_join h _ ?. exfalso.
+        hauto l:on use:REReds.hne_app_inv, REReds.var_inv.
       (* real case *)
       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
@@ -753,11 +756,14 @@ Proof.
         have : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
         hauto l:on use:bind_inst.
         hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
-      * admit.
+      * move => p p0 p1 p2 /algo_metric_join. clear => ? ? ?. exfalso.
+        hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
     + case : b => //=.
-      * admit.
-      * admit.
+      * move => i p h /algo_metric_join. clear => ? _ ?. exfalso.
+        hauto l:on use:REReds.hne_proj_inv, REReds.var_inv.
+      * move => > /algo_metric_join. clear => ? ? ?. exfalso.
+        hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv.
       (* real case *)
       * admit.
       * sfirstorder use:T_Bot_Imp.

From 9bd554b513d542a1bc546514b37a8fdee7322c5f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 14 Feb 2025 21:55:44 -0500
Subject: [PATCH 100/210] Add injectivity lemma for abs

---
 theories/algorithmic.v |  1 +
 theories/fp_red.v      | 33 +++++++++++++++++++++++++++++++++
 2 files changed, 34 insertions(+)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 5aa0a3b..1496d4c 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -681,6 +681,7 @@ Proof.
       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
       repeat split =>//=. sfirstorder.
       simpl in *; by lia.
+
       admit.
     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 50d8491..b42033c 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2319,6 +2319,14 @@ Module DJoin.
     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).
+  Proof. substify. apply substing. Qed.
+
+  Lemma weakening n (a b : PTm n)  :
+    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) :
     R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) a).
   Proof.
@@ -2327,6 +2335,31 @@ Module DJoin.
     hauto q:on ctrs:rtc inv:option use:REReds.cong.
   Qed.
 
+  Lemma abs_inj n (a0 a1 : PTm (S n)) :
+    SN a0 -> SN a1 ->
+    R (PAbs a0) (PAbs a1) ->
+    R a0 a1.
+  Proof.
+    move => sn0 sn1.
+    move /weakening => /=.
+    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 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.
+    apply : transitive; try apply h1.
+    apply : N_Exp. apply N_β. sauto.
+    by asimpl.
+    eauto.
+  Qed.
+
 End DJoin.
 
 

From 67f91970d624cdff8ea321290d26be6bc92d3dc8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 15 Feb 2025 14:04:04 -0500
Subject: [PATCH 101/210] Finish the admitted inversion lemmas that depend on
 SN

---
 theories/algorithmic.v | 153 +++++++++++++++++++++++++++++++++++------
 theories/fp_red.v      |   6 ++
 2 files changed, 139 insertions(+), 20 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 1496d4c..456f1f4 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -59,23 +59,143 @@ Lemma E_Conv_E n Γ (a b : PTm n) A B i :
   Γ ⊢ a ≡ b ∈ B.
 Proof. qauto use:E_Conv, Su_Eq, E_Symmetric. Qed.
 
+Lemma lored_embed n Γ (a b : PTm n) 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 :
+  Γ ⊢ a ∈ A -> rtc LoRed.R a b -> Γ ⊢ a ≡ b ∈ A.
+Proof.
+  move => + h. move : Γ A.
+  elim : a b /h.
+  - sfirstorder use:E_Refl.
+  - move => a b c ha hb ih Γ A hA.
+    have ? : Γ ⊢ a ≡ b ∈ A by sfirstorder use:lored_embed.
+    have ? : Γ ⊢ b ∈ A by hauto l:on use:regularity.
+    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 Sub_Bind_InvR n Γ p (A : PTm n) B C :
   Γ ⊢ PBind p A B ≲ C ->
   exists i A0 B0, Γ ⊢ C ≡ PBind p A0 B0 ∈ PUniv i.
 Proof.
-Admitted.
+  move => /[dup] h.
+  move /synsub_to_usub.
+  move => [h0 [h1 h2]].
+  move /LoReds.FromSN : h1.
+  move => [C' [hC0 hC1]].
+  have [i hC] : exists i, Γ ⊢ C ∈ PUniv i by qauto l:on use:regularity.
+  have hE : Γ ⊢ C ≡ C' ∈ PUniv i by sfirstorder use:loreds_embed.
+  have : Γ ⊢ PBind p A B ≲ C' by eauto using Su_Transitive, Su_Eq.
+  move : hE hC1. clear.
+  case : C' => //=.
+  - move => k _ _ /synsub_to_usub.
+    clear. move => [_ [_ h]]. exfalso.
+    rewrite /Sub.R in h.
+    move : h => [c][d][+ []].
+    move /REReds.bind_inv => ?.
+    move /REReds.var_inv => ?.
+    sauto lq:on.
+  - move => p0 h. exfalso.
+    have {h} : Γ ⊢ PAbs p0 ∈ PUniv i by hauto l:on use:regularity.
+    move /Abs_Inv => [A0][B0][_]/synsub_to_usub.
+    hauto l:on use:Sub.bind_univ_noconf.
+  - move => u v hC /andP [h0 h1] /synsub_to_usub ?.
+    exfalso.
+    suff : SNe (PApp u v) by hauto l:on use:Sub.bind_sne_noconf.
+    eapply ne_nf_embed => //=. itauto.
+  - move => p0 p1 hC  h  ?. exfalso.
+    have {hC} : Γ ⊢ PPair p0 p1 ∈ PUniv i by hauto l:on use:regularity.
+    hauto lq:on use:Sub.bind_univ_noconf, synsub_to_usub, Pair_Inv.
+  - move => p0 p1 _ + /synsub_to_usub.
+    hauto lq:on use:Sub.bind_sne_noconf, ne_nf_embed.
+  - move => b p0 p1 h0 h1 /[dup] h2 /synsub_to_usub *.
+    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.
+Qed.
+
+Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
+  Γ ⊢ PUniv i ≲ C ->
+  exists j k, Γ ⊢ C ≡ PUniv j ∈ PUniv k.
+Proof.
+  move => /[dup] h.
+  move /synsub_to_usub.
+  move => [h0 [h1 h2]].
+  move /LoReds.FromSN : h1.
+  move => [C' [hC0 hC1]].
+  have [j hC] : exists i, Γ ⊢ C ∈ PUniv i by qauto l:on use:regularity.
+  have hE : Γ ⊢ C ≡ C' ∈ PUniv j by sfirstorder use:loreds_embed.
+  have : Γ ⊢ PUniv i ≲ C' by eauto using Su_Transitive, Su_Eq.
+  move : hE hC1. clear.
+  case : C' => //=.
+  - move => f => _ _ /synsub_to_usub.
+    move => [_ [_]].
+    move => [v0][v1][/REReds.univ_inv + [/REReds.var_inv ]].
+    hauto lq:on inv:Sub1.R.
+  - move => p hC.
+    have {hC} : Γ ⊢ PAbs p ∈ PUniv j by hauto l:on use:regularity.
+    hauto lq:on rew:off use:Abs_Inv, synsub_to_usub,
+            Sub.bind_univ_noconf.
+  - hauto q:on use:synsub_to_usub, Sub.univ_sne_noconf, ne_nf_embed.
+  - move => a b hC.
+    have {hC} : Γ ⊢ PPair a b ∈ PUniv j by hauto l:on use:regularity.
+    hauto lq:on rew:off use:Pair_Inv, synsub_to_usub,
+            Sub.bind_univ_noconf.
+  - 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.
+Qed.
 
 Lemma Sub_Bind_InvL n Γ p (A : PTm n) B C :
   Γ ⊢ C ≲ PBind p A B ->
   exists i A0 B0, Γ ⊢ PBind p A0 B0 ≡ C ∈ PUniv i.
 Proof.
-Admitted.
-
-Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
-  Γ ⊢ PUniv i ≲ C ->
-  exists j, Γ ⊢ C ≡ PUniv j ∈ PUniv (S j).
-Proof.
-Admitted.
+  move => /[dup] h.
+  move /synsub_to_usub.
+  move => [h0 [h1 h2]].
+  move /LoReds.FromSN : h0.
+  move => [C' [hC0 hC1]].
+  have [i hC] : exists i, Γ ⊢ C ∈ PUniv i by qauto l:on use:regularity.
+  have hE : Γ ⊢ C ≡ C' ∈ PUniv i by sfirstorder use:loreds_embed.
+  have : Γ ⊢ C' ≲ PBind p A B by eauto using Su_Transitive, Su_Eq, E_Symmetric.
+  move : hE hC1. clear.
+  case : C' => //=.
+  - move => k _ _ /synsub_to_usub.
+    clear. move => [_ [_ h]]. exfalso.
+    rewrite /Sub.R in h.
+    move : h => [c][d][+ []].
+    move /REReds.var_inv => ?.
+    move /REReds.bind_inv => ?.
+    hauto lq:on inv:Sub1.R.
+  - move => p0 h. exfalso.
+    have {h} : Γ ⊢ PAbs p0 ∈ PUniv i by hauto l:on use:regularity.
+    move /Abs_Inv => [A0][B0][_]/synsub_to_usub.
+    hauto l:on use:Sub.bind_univ_noconf.
+  - move => u v hC /andP [h0 h1] /synsub_to_usub ?.
+    exfalso.
+    suff : SNe (PApp u v) by hauto l:on use:Sub.sne_bind_noconf.
+    eapply ne_nf_embed => //=. itauto.
+  - move => p0 p1 hC  h  ?. exfalso.
+    have {hC} : Γ ⊢ PPair p0 p1 ∈ PUniv i by hauto l:on use:regularity.
+    hauto lq:on use:Sub.bind_univ_noconf, synsub_to_usub, Pair_Inv.
+  - move => p0 p1 _ + /synsub_to_usub.
+    hauto lq:on use:Sub.sne_bind_noconf, ne_nf_embed.
+  - move => b p0 p1 h0 h1 /[dup] h2 /synsub_to_usub *.
+    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.
+Qed.
 
 Lemma T_Abs_Inv n Γ (a0 a1 : PTm (S n)) U :
   Γ ⊢ PAbs a0 ∈ U ->
@@ -210,6 +330,7 @@ Lemma coqeq_symmetric_mutual : forall n,
     (forall (a b : PTm n), a ⇔ b -> b ⇔ a).
 Proof. apply coqeq_mutual; qauto l:on ctrs:CoqEq,CoqEq_R, CoqEq_Neu. Qed.
 
+
 (* Lemma Sub_Univ_InvR *)
 
 Lemma coqeq_sound_mutual : forall n,
@@ -496,7 +617,7 @@ Lemma T_PairBind_Imp n Γ (a0 a1 : PTm n) p A0 B0 U :
 Proof.
   move /Pair_Inv => [A1][B1][_][_]hbU.
   move /Bind_Inv => [i][_ [_ haU]].
-  move /Sub_Univ_InvR : haU => [j]hU.
+  move /Sub_Univ_InvR : haU => [j][k]hU.
   have : Γ ⊢ PBind PSig A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_univ_noconf.
 Qed.
@@ -518,7 +639,7 @@ Lemma T_PairUniv_Imp n Γ (a0 a1 : PTm n) i U :
 Proof.
   move /Pair_Inv => [A1][B1][_][_]hbU.
   move /Univ_Inv => [h0 h1].
-  move /Sub_Univ_InvR : h1 => [j hU].
+  move /Sub_Univ_InvR : h1 => [j [k hU]].
   have : Γ ⊢ PBind PSig A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
   clear. move /synsub_to_usub.
   hauto lq:on use:Sub.bind_univ_noconf.
@@ -531,7 +652,7 @@ Lemma T_AbsUniv_Imp n Γ a i (A : PTm n) :
 Proof.
   move /Abs_Inv => [A0][B0][_]haU.
   move /Univ_Inv => [h0 h1].
-  move /Sub_Univ_InvR : h1 => [j hU].
+  move /Sub_Univ_InvR : h1 => [j [k hU]].
   have : Γ ⊢ PBind PPi A0 B0 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
   clear. move /synsub_to_usub.
   hauto lq:on use:Sub.bind_univ_noconf.
@@ -544,19 +665,12 @@ Lemma T_AbsBind_Imp n Γ a p A0 B0 (U : PTm n) :
 Proof.
   move /Abs_Inv => [A1][B1][_ ha].
   move /Bind_Inv => [i [_ [_ h]]].
-  move /Sub_Univ_InvR : h => [j hU].
+  move /Sub_Univ_InvR : h => [j [k hU]].
   have : Γ ⊢ PBind PPi A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
   clear. move /synsub_to_usub.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
-Lemma T_Bot_Imp n Γ (A : PTm n) :
-  Γ ⊢ PBot ∈ A -> False.
-  move E : PBot => u hu.
-  move : E.
-  induction hu => //=.
-Qed.
-
 Lemma lored_nsteps_abs_inv k n (a : PTm (S n)) b :
   nsteps LoRed.R k (PAbs a) b -> exists b', nsteps LoRed.R k a b' /\ b = PAbs b'.
 Proof.
@@ -681,7 +795,6 @@ Proof.
       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
       repeat split =>//=. sfirstorder.
       simpl in *; by lia.
-
       admit.
     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index b42033c..536c561 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -293,6 +293,12 @@ Proof.
   apply N_β'. by asimpl. eauto.
 Qed.
 
+Lemma ne_nf_embed n (a : PTm n) :
+  (ne a -> SNe a) /\ (nf a -> SN a).
+Proof.
+  elim : n  / a => //=; hauto qb:on ctrs:SNe, SN.
+Qed.
+
 #[export]Hint Constructors SN SNe TRedSN : sn.
 
 Ltac2 rec solve_anti_ren () :=

From 9d951a24c56bf2a9eb98d5f6469ae2a9cef778eb Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 15 Feb 2025 14:31:31 -0500
Subject: [PATCH 102/210] Add standardization theorem for djoin

---
 theories/fp_red.v | 20 ++++++++++++++++++++
 1 file changed, 20 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 536c561..f7ca253 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2366,6 +2366,26 @@ Module DJoin.
     eauto.
   Qed.
 
+  Lemma standardization n (a b : PTm n) :
+    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.
+    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.
+    move : h0 h1 hv3. clear.
+    move => h0 h1 hv hbv hav.
+    move : rered_standardization (h0) hav. repeat move/[apply].
+    move => [a1 [ha0 ha1]].
+    move : rered_standardization (h1) hbv. repeat move/[apply].
+    move => [b1 [hb0 hb1]].
+    have [*] : nf a1 /\ nf b1 by sfirstorder use:NeEPars.R_nonelim_nf.
+    hauto q:on use:NeEPars.ToEReds unfold:EJoin.R.
+  Qed.
+
 End DJoin.
 
 

From 926c2284a5feb981901d0d9dbfddc4569afc5c12 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 15 Feb 2025 16:39:05 -0500
Subject: [PATCH 103/210] Finish the pair pair case

---
 theories/algorithmic.v | 154 ++++++++++++++++++++++++++++++++++++++---
 theories/fp_red.v      |   7 ++
 2 files changed, 150 insertions(+), 11 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 456f1f4..263293b 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -709,6 +709,47 @@ 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) :
+    nsteps LoRed.R k (PProj p a0) C ->
+    ishne a0 ->
+    exists i a1,
+      i <= k /\
+        C = PProj p a1 /\
+        nsteps LoRed.R i a0 a1.
+Proof.
+  move E : (PProj p a0) => u hu. move : a0 E.
+  elim : k u C / hu.
+  - sauto lq:on.
+  - move => k a0 a1 a2 ha01 ha12 ih a3 ?. subst.
+    inversion ha01; subst => //=.
+    spec_refl.
+    move => h.
+    have : ishne a4 by sfirstorder use:lored_hne_preservation.
+    move : ih => /[apply]. move => [i][a1][?][?]h0. subst.
+    exists (S i), a1. hauto lq:on ctrs:nsteps solve+:lia.
+Qed.
+
+Lemma algo_metric_proj n k p0 p1 (a0 a1 : PTm n) :
+  algo_metric k (PProj p0 a0) (PProj p1 a1) ->
+  ishne a0 ->
+  ishne a1 ->
+  p0 = p1 /\ exists j, j < k /\ algo_metric j a0 a1.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 hne0 hne1.
+  move : lored_nsteps_proj_inv h0 (hne0);repeat move/[apply].
+  move => [i0][a2][hi][?]ha02. subst.
+  move : lored_nsteps_proj_inv h1 (hne1);repeat move/[apply].
+  move => [i1][a3][hj][?]ha13. subst.
+  simpl in *.
+  move /EJoin.hne_proj_inj : h4 => [h40 h41]. subst.
+  split => //.
+  exists (k - 1). split. simpl in *; lia.
+  rewrite/algo_metric.
+  do 4 eexists. repeat split; eauto. sfirstorder use:ne_nf.
+  sfirstorder use:ne_nf.
+  lia.
+Qed.
+
 Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
   algo_metric k (PApp a0 b0) (PApp a1 b1) ->
   ishne a0 ->
@@ -754,7 +795,7 @@ Qed.
 Lemma coqeq_complete n k (a b : PTm n) :
   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).
+  (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.
   elim /Wf_nat.lt_wf_ind.
@@ -786,7 +827,7 @@ Proof.
       suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto.
       have ? : n > 0 by sauto solve+:lia.
       exists (n - 1). split; first by lia.
-      move : ha0. rewrite /algo_metric.
+      move : (ha0). rewrite /algo_metric.
       move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har.
       apply lored_nsteps_abs_inv in hr0, hr1.
       move : hr0 => [va' [hr00 hr01]].
@@ -795,7 +836,19 @@ Proof.
       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
       repeat split =>//=. sfirstorder.
       simpl in *; by lia.
-      admit.
+      move /algo_metric_join /DJoin.symmetric : ha0.
+      have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
+      move => /[dup] [[ha00 ha10]] [].
+      move : DJoin.abs_inj; repeat move/[apply].
+      move : DJoin.standardization ha00 ha10; repeat move/[apply].
+      move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]].
+      have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va'
+        by hauto lq:on use:@relations.rtc_nsteps.
+      have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb'
+        by hauto lq:on use:@relations.rtc_nsteps.
+      simpl in *.
+      have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst.
+      sfirstorder.
     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
       apply : CE_HRed; eauto using rtc_refl.
@@ -804,17 +857,34 @@ Proof.
       (* move => [l [hl [hal0 hal1]]]. *)
       (* apply CE_PairPair. eapply ih; eauto. *)
       (* by eapply ih; eauto. *)
-    + admit.
-    + admit.
+    + case : b => //=.
+      * hauto lq:on use:T_AbsBind_Imp.
+      * hauto lq:on rew:off use:T_PairBind_Imp.
+      * move => p1 A1 B1 p0 A0 B0.
+      (* Show that A0 and A1 are algorithmically equal *)
+      (* Use soundness to show that they are actually definitionally equal *)
+      (* Use that to show that B0 and B1 can be assigned the same type *)
+        admit.
+      * move => > /algo_metric_join.
+        hauto lq:on use:DJoin.bind_univ_noconf.
+    + case : b => //=.
+      * hauto lq:on use:T_AbsUniv_Imp.
+      * hauto lq:on use:T_PairUniv_Imp.
+      * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric.
+      * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj.
+        subst.
+        hauto l:on.
   - move : k a b h fb fa. abstract : hnfneu.
-    admit.
+    move => k.
+    move => + b.
+    case : b => //=; admit.
   - 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)).
+    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)).
     move => h0.
     split. move => *. apply : CE_HRed; eauto using rtc_refl. apply CE_NeuNeu. by firstorder.
     by firstorder.
@@ -825,7 +895,7 @@ Proof.
       * have ? : j = i by hauto lq:on use:algo_metric_join, DJoin.var_inj. subst.
         move => Γ A B hA hB.
         split. apply CE_VarCong.
-        exists (Γ i). hauto l:on use:Var_Inv.
+        exists (Γ i). hauto l:on use:Var_Inv, T_Var.
       * move => p p0 f /algo_metric_join. clear => ? ? _. exfalso.
         hauto l:on use:REReds.var_inv, REReds.hne_app_inv.
       * move => a0 a1 i /algo_metric_join. clear => ? ? _. exfalso.
@@ -844,7 +914,7 @@ Proof.
         move /(_ hj).
         move => [_ ihb].
         move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply].
-        move => [hb01][C][hT0]hT1.
+        move => [hb01][C][hT0][hT1][hT2]hT3.
         move /Sub_Bind_InvL : (hT0).
         move => [i][A2][B2]hE.
         have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by
@@ -865,11 +935,21 @@ Proof.
         apply : Su_Transitive; eauto.
         move /E_Refl in ha0.
         hauto l:on use:Su_Pi_Proj2.
+        move => [:hsub].
+        have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
+        split.
         apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
         move /regularity_sub0 : hSu10 => [i0].
-        have : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
         hauto l:on use:bind_inst.
         hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
+        split.
+        by apply : T_App; eauto using T_Conv_E.
+        apply : T_Conv; eauto.
+        apply T_App with (A := A2) (B := B2); eauto.
+        apply : T_Conv_E; eauto.
+        move /E_Symmetric in h01.
+        move /regularity_sub0 : hSu20 => [i0].
+        sfirstorder use:bind_inst.
       * move => p p0 p1 p2 /algo_metric_join. clear => ? ? ?. exfalso.
         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
@@ -879,7 +959,59 @@ Proof.
       * move => > /algo_metric_join. clear => ? ? ?. exfalso.
         hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv.
       (* real case *)
-      * admit.
+      * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0.
+        move : ha (hne0) (hne1); repeat move/[apply].
+        move => [? [j []]]. subst.
+        move : ih; repeat move/[apply].
+        move => [_ ih].
+        case : p1.
+        ** move => Γ A B ha0 ha1.
+           move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0.
+           move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1.
+           move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply].
+           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
+           split. sauto lq:on.
+           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
+           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
+             by eauto using Su_Transitive, Su_Eq.
+           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
+             by eauto using Su_Transitive, Su_Eq.
+           exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive.
+           repeat split => //=.
+           hauto l:on use:Su_Sig_Proj1, Su_Transitive.
+           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
+           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
+        ** move => Γ A B  ha0  ha1.
+           move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0.
+           move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1.
+           move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply].
+           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
+           split. sauto lq:on.
+           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
+           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
+             by eauto using Su_Transitive, Su_Eq.
+           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
+             by eauto using Su_Transitive, Su_Eq.
+           have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1.
+           have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1.
+           have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+           have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+           have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2
+             by sauto lq:on use:coqeq_sound_mutual.
+           exists (subst_PTm (scons (PProj PL a0) VarPTm) B2).
+           repeat split.
+           *** apply : Su_Transitive; eauto.
+               have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2
+                 by qauto use:regularity, E_Refl.
+               sfirstorder use:Su_Sig_Proj2.
+           *** apply : Su_Transitive; eauto.
+               sfirstorder use:Su_Sig_Proj2.
+           *** eauto using T_Proj2.
+           *** apply : T_Conv.
+               apply : T_Proj2; eauto.
+               move /E_Symmetric in haE.
+               move /regularity_sub0 in hSu21.
+               sfirstorder use:bind_inst.
       * sfirstorder use:T_Bot_Imp.
     + sfirstorder use:T_Bot_Imp.
 Admitted.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index f7ca253..e5b2b54 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1554,6 +1554,13 @@ Module EJoin.
     hauto lq:on use:EReds.app_inv.
   Qed.
 
+  Lemma hne_proj_inj n p0 p1 (a0 a1 : PTm n) :
+    R (PProj p0 a0) (PProj p1 a1) ->
+    p0 = p1 /\ R a0 a1.
+  Proof.
+    hauto lq:on rew:off use:EReds.proj_inv.
+  Qed.
+
 End EJoin.
 
 Module RERed.

From 3fb6d411e793e92000e3e2a66b625ac90b3d4885 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 15 Feb 2025 17:22:43 -0500
Subject: [PATCH 104/210] Finish most of the pi pi case

---
 theories/algorithmic.v | 72 ++++++++++++++++++++++++++++++++++++++++--
 theories/fp_red.v      | 18 +++++++++++
 2 files changed, 88 insertions(+), 2 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 263293b..27edc21 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -685,6 +685,26 @@ Lemma lored_hne_preservation n (a b : PTm n) :
     LoRed.R a b -> ishne a -> ishne b.
 Proof.  induction 1; sfirstorder. Qed.
 
+Lemma lored_nsteps_bind_inv k n p (a0 : PTm n) b0 C :
+    nsteps LoRed.R k (PBind p a0 b0) C ->
+    exists i j a1 b1,
+      i <= k /\ j <= k /\
+        C = PBind p a1 b1 /\
+        nsteps LoRed.R i a0 a1 /\
+        nsteps LoRed.R j b0 b1.
+Proof.
+  move E : (PBind p a0 b0) => u hu. move : p a0 b0 E.
+  elim : k u C / hu.
+  - sauto lq:on.
+  - move => k a0 a1 a2 ha ha' ih p a3 b0 ?. subst.
+    inversion ha; subst => //=;
+    spec_refl.
+    move : ih => [i][j][a1][b1][?][?][?][h0]h1. subst.
+    exists (S i),j,a1,b1. sauto lq:on solve+:lia.
+    move : ih => [i][j][a1][b1][?][?][?][h0]h1. subst.
+    exists i,(S j),a1,b1. sauto lq:on solve+:lia.
+Qed.
+
 Lemma lored_nsteps_app_inv k n (a0 b0 C : PTm n) :
     nsteps LoRed.R k (PApp a0 b0) C ->
     ishne a0 ->
@@ -779,6 +799,25 @@ Proof.
     repeat split => //=; sfirstorder b:on use:ne_nf.
 Qed.
 
+Lemma algo_metric_bind n k p0 (A0 : PTm n) 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.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
+  move : lored_nsteps_bind_inv h0 => /[apply].
+  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
+  move : lored_nsteps_bind_inv h1 => /[apply].
+  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+  move /EJoin.bind_inj : h4 => [?][hEa]hEb. subst.
+  split => //.
+  exists (k - 1). split. simpl in *; lia.
+  simpl in *.
+  split.
+  - exists i0,j0,a2,a3.
+    repeat split => //=; sfirstorder b:on solve+:lia.
+  - exists i1,j1,b2,b3. sfirstorder b:on solve+:lia.
+Qed.
+
 Lemma algo_metric_join n k (a b : PTm n) :
   algo_metric k a b ->
   DJoin.R a b.
@@ -861,10 +900,31 @@ Proof.
       * hauto lq:on use:T_AbsBind_Imp.
       * hauto lq:on rew:off use:T_PairBind_Imp.
       * move => p1 A1 B1 p0 A0 B0.
+        move /algo_metric_bind.
+        move => [?][j [ih0 [ih1 ih2]]]_ _. subst.
+        move => Γ A hU0 hU1.
+        move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]].
+        move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]].
+        have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1)
+          by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia.
+        have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1)
+          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 ihA : A0 ⇔ A1 by hauto l:on.
+        have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1)
+          by hauto l:on use:coqeq_sound_mutual.
+        apply : CE_HRed; eauto using rtc_refl.
+        constructor => //.
+        admit.
+        (* eapply ih; eauto. *)
+        (* move /Su_Eq in hAE. *)
+        (* apply : ctx_eq_subst_one; eauto. *)
+
       (* Show that A0 and A1 are algorithmically equal *)
       (* Use soundness to show that they are actually definitionally equal *)
       (* Use that to show that B0 and B1 can be assigned the same type *)
-        admit.
+        (* admit. *)
       * move => > /algo_metric_join.
         hauto lq:on use:DJoin.bind_univ_noconf.
     + case : b => //=.
@@ -877,7 +937,15 @@ Proof.
   - move : k a b h fb fa. abstract : hnfneu.
     move => k.
     move => + b.
-    case : b => //=; admit.
+    case : b => //=.
+    (* NeuAbs *)
+    + admit.
+    (* NeuPair *)
+    + admit.
+    (* NeuBind: Impossible *)
+    + admit.
+    (* NeuUniv: Impossible *)
+    + admit.
   - move {ih}.
     move /algo_metric_sym in h.
     qauto l:on use:coqeq_symmetric_mutual.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index e5b2b54..31789ba 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1540,6 +1540,17 @@ Module EReds.
       hauto q:on ctrs:rtc,ERed.R inv:ERed.R.
   Qed.
 
+  Lemma bind_inv n p (A : PTm n) 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.
+    move E : (PBind p A B) => u hu.
+    move : p A B E.
+    elim : u C / hu.
+    hauto lq:on ctrs:rtc.
+    hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
+  Qed.
+
 End EReds.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -1561,6 +1572,13 @@ Module EJoin.
     hauto lq:on rew:off use:EReds.proj_inv.
   Qed.
 
+  Lemma bind_inj n p0 p1 (A0 A1 : PTm n) 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.
+
 End EJoin.
 
 Module RERed.

From 0f48a485be965faae7708a157bdf45309624a4e1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 01:13:41 -0500
Subject: [PATCH 105/210] Prove some impossible cases

---
 theories/algorithmic.v | 35 +++++++++++++++++++++++------------
 theories/fp_red.v      | 26 ++++++++++++++++++++++++++
 2 files changed, 49 insertions(+), 12 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 27edc21..1ba1e80 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -831,6 +831,12 @@ Lemma algo_metric_join n k (a b : PTm n) :
   rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
 Qed.
 
+Lemma T_Univ_Raise n Γ (a : PTm n) 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 coqeq_complete n k (a b : PTm n) :
   algo_metric k a b ->
   (forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) /\
@@ -909,22 +915,24 @@ Proof.
           by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia.
         have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1)
           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 {}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)
+          by hauto lq:on use:T_Univ_Raise solve+:lia.
+        have {}hB1 : funcomp (ren_PTm shift) (scons A1 Γ) ⊢
+                       B1 ∈ PUniv (max i0 i1)
+          by hauto lq:on use:T_Univ_Raise solve+:lia.
+
         have ihA : A0 ⇔ A1 by hauto l:on.
         have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1)
           by hauto l:on use:coqeq_sound_mutual.
         apply : CE_HRed; eauto using rtc_refl.
         constructor => //.
-        admit.
-        (* eapply ih; eauto. *)
-        (* move /Su_Eq in hAE. *)
-        (* apply : ctx_eq_subst_one; eauto. *)
 
-      (* Show that A0 and A1 are algorithmically equal *)
-      (* Use soundness to show that they are actually definitionally equal *)
-      (* Use that to show that B0 and B1 can be assigned the same type *)
-        (* admit. *)
+        eapply ih; eauto.
+        apply : ctx_eq_subst_one; eauto.
+        eauto using Su_Eq.
       * move => > /algo_metric_join.
         hauto lq:on use:DJoin.bind_univ_noconf.
     + case : b => //=.
@@ -943,9 +951,12 @@ Proof.
     (* NeuPair *)
     + admit.
     (* NeuBind: Impossible *)
-    + admit.
+    + move => b p p0 a /algo_metric_join h _ h0. exfalso.
+      move : h h0. clear.
+      move /Sub.FromJoin.
+      hauto l:on use:Sub.hne_bind_noconf.
     (* NeuUniv: Impossible *)
-    + admit.
+    + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join.
   - move {ih}.
     move /algo_metric_sym in h.
     qauto l:on use:coqeq_symmetric_mutual.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 31789ba..e00994c 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1675,6 +1675,10 @@ Module RERed.
 End RERed.
 
 Module REReds.
+  Lemma hne_preservation n (a b : PTm n) :
+    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) :
     rtc RERed.R a b ->
     SN a ->
@@ -2284,6 +2288,17 @@ Module DJoin.
     case : c => //=; itauto.
   Qed.
 
+  Lemma hne_univ_noconf n (a b : PTm n) :
+    R a b -> ishne a -> isuniv b -> False.
+  Proof.
+    move => [c [h0 h1]] h2 h3.
+    have {h0 h1 h2 h3} : ishne c /\ isuniv c by
+      hauto l:on use:REReds.hne_preservation,
+          REReds.univ_preservation.
+    move => [].
+    case : c => //=.
+  Qed.
+
   Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
     DJoin.R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ DJoin.R A0 A1 /\ DJoin.R B0 B1.
@@ -2564,6 +2579,17 @@ Module Sub.
     qauto l:on inv:SNe.
   Qed.
 
+  Lemma hne_bind_noconf n (a b : PTm n) :
+    R a b -> ishne a -> isbind b -> False.
+  Proof.
+    rewrite /R.
+    move => [c[d] [? []]] h0 h1 h2 h3.
+    have : ishne c by eauto using REReds.hne_preservation.
+    have : isbind d by eauto using REReds.bind_preservation.
+    move : h1. clear. inversion 1; subst => //=.
+    clear. case : d => //=.
+  Qed.
+
   Lemma bind_sne_noconf n (a b : PTm n) :
     R a b -> SNe b -> isbind a -> False.
   Proof.

From 06d420aa7ee724db4d3298d60282b456e51b3358 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 01:22:15 -0500
Subject: [PATCH 106/210] Add stubs for lemmas needed for completeness

---
 theories/algorithmic.v | 18 ++++++++++++++++++
 1 file changed, 18 insertions(+)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 1ba1e80..460d149 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -558,6 +558,7 @@ Proof.
   - sfirstorder use:ne_hne.
   - hauto lq:on ctrs:HRed.R.
 Qed.
+
 Lemma algo_metric_case n k (a b : PTm n) :
   algo_metric k a b ->
   (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ algo_metric k' a' b /\ k' < k.
@@ -770,6 +771,23 @@ Proof.
   lia.
 Qed.
 
+Lemma algo_metric_pair n k (a0 b0 a1 b1 : PTm n) :
+  SN (PPair a0 b0) ->
+  SN (PPair a1 b1) ->
+  algo_metric k (PPair a0 b0) (PPair a1 b1) ->
+  exists j, j < k /\ algo_metric j a0 a1 /\ algo_metric j b0 b1.
+Admitted.
+
+Lemma algo_metric_neu_abs n k (a0 : PTm (S n)) u :
+  algo_metric k u (PAbs a0) ->
+  exists j, j < k /\ algo_metric j (PApp (ren_PTm shift u) (VarPTm var_zero)) a0.
+Admitted.
+
+Lemma algo_metric_neu_pair n k (a0 b0 : PTm n) u :
+  algo_metric k u (PPair a0 b0) ->
+  exists j, j < k /\ algo_metric j (PProj PL u) a0 /\ algo_metric j (PProj PR u) b0.
+Admitted.
+
 Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
   algo_metric k (PApp a0 b0) (PApp a1 b1) ->
   ishne a0 ->

From 60a4eb886fd7e3982964870733f2b729178b205a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 19:14:54 -0500
Subject: [PATCH 107/210] Finish injectivity for pairs

---
 theories/algorithmic.v | 59 ++++++++++++++++++++++++++--------
 theories/fp_red.v      | 73 ++++++++++++++++++++++++++++++++++++++----
 2 files changed, 112 insertions(+), 20 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 460d149..24ce2b9 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -771,12 +771,56 @@ Proof.
   lia.
 Qed.
 
+Lemma lored_nsteps_pair_inv k n (a0 b0 C : PTm n) :
+    nsteps LoRed.R k (PPair a0 b0) C ->
+    exists i j a1 b1,
+      i <= k /\ j <= k /\
+        C = PPair a1 b1 /\
+        nsteps LoRed.R i a0 a1 /\
+        nsteps LoRed.R j b0 b1.
+  move E : (PPair a0 b0) => u hu. move : a0 b0 E.
+  elim : k u C / hu.
+  - sauto lq:on.
+  - move => k a0 a1 a2 ha01 ha12 ih a3 b0 ?.  subst.
+    inversion ha01; subst => //=.
+    spec_refl.
+    move : ih => [i][j][a1][b1][?][?][?][h0]h1.
+    subst. exists (S i),j,a1,b1. sauto lq:on solve+:lia.
+    spec_refl.
+    move : ih => [i][j][a1][b1][?][?][?][h0]h1. subst.
+    exists i, (S j), a1, b1. sauto lq:on solve+:lia.
+Qed.
+
+Lemma algo_metric_join n k (a b : PTm n) :
+  algo_metric k a b ->
+  DJoin.R a b.
+  rewrite /algo_metric.
+  move => [i][j][va][vb][h0][h1][h2][h3][[v [h40 h41]]]h5.
+  have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
+  have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
+  apply REReds.FromEReds in h40,h41.
+  apply LoReds.ToRReds in h0,h1.
+  apply REReds.FromRReds in h0,h1.
+  rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
+Qed.
+
 Lemma algo_metric_pair n k (a0 b0 a1 b1 : PTm n) :
   SN (PPair a0 b0) ->
   SN (PPair a1 b1) ->
   algo_metric k (PPair a0 b0) (PPair a1 b1) ->
   exists j, j < k /\ algo_metric j a0 a1 /\ algo_metric j b0 b1.
-Admitted.
+  move => sn0 sn1 /[dup] /algo_metric_join hj.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
+  move : lored_nsteps_pair_inv h0;repeat move/[apply].
+  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
+  move : lored_nsteps_pair_inv h1;repeat move/[apply].
+  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+  simpl in *. exists (k - 1).
+  move /andP : h2 => [h20 h21].
+  move /andP : h3 => [h30 h31].
+  suff : EJoin.R a2 a3 /\ EJoin.R b2 b3 by hauto lq:on solve+:lia.
+  hauto l:on use:DJoin.ejoin_pair_inj.
+Qed.
 
 Lemma algo_metric_neu_abs n k (a0 : PTm (S n)) u :
   algo_metric k u (PAbs a0) ->
@@ -836,19 +880,6 @@ Proof.
   - exists i1,j1,b2,b3. sfirstorder b:on solve+:lia.
 Qed.
 
-Lemma algo_metric_join n k (a b : PTm n) :
-  algo_metric k a b ->
-  DJoin.R a b.
-  rewrite /algo_metric.
-  move => [i][j][va][vb][h0][h1][h2][h3][[v [h40 h41]]]h5.
-  have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
-  have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
-  apply REReds.FromEReds in h40,h41.
-  apply LoReds.ToRReds in h0,h1.
-  apply REReds.FromRReds in h0,h1.
-  rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
-Qed.
-
 Lemma T_Univ_Raise n Γ (a : PTm n) i j :
   Γ ⊢ a ∈ PUniv i ->
   i <= j ->
diff --git a/theories/fp_red.v b/theories/fp_red.v
index e00994c..407d3f5 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -262,6 +262,12 @@ end.
 Lemma ne_nf n a : @ne n a -> nf a.
 Proof. elim : a => //=. Qed.
 
+Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
+  (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)).
+Proof.
+  move : m ξ. elim : n / a => //=; solve [hauto b:on].
+Qed.
+
 Inductive TRedSN' {n} (a : PTm n) : PTm n -> Prop :=
 | T_Refl :
   TRedSN' a a
@@ -842,12 +848,6 @@ Module RReds.
 End RReds.
 
 
-Lemma ne_nf_ren n m (a : PTm n) (ξ : fin n -> fin m) :
-  (ne a <-> ne (ren_PTm ξ a)) /\ (nf a <-> nf (ren_PTm ξ a)).
-Proof.
-  move : m ξ. elim : n / a => //=; solve [hauto b:on].
-Qed.
-
 Module NeEPar.
   Inductive R_nonelim {n} : PTm n -> PTm n ->  Prop :=
   (****************** Eta ***********************)
@@ -1446,6 +1446,14 @@ Module ERed.
     all : hauto lq:on ctrs:R.
   Qed.
 
+  Lemma nf_preservation n (a b : PTm n) :
+    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.
+  Qed.
+
 End ERed.
 
 Module EReds.
@@ -1829,6 +1837,14 @@ Module REReds.
     hauto l:on use:substing.
   Qed.
 
+  Lemma ToEReds n (a b : PTm n) :
+    nf a ->
+    rtc RERed.R a b -> rtc ERed.R a b.
+  Proof.
+    move => + h.
+    induction h; hauto lq:on rew:off ctrs:rtc use:RERed.ToBetaEta, RReds.nf_refl, @rtc_once, ERed.nf_preservation.
+  Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2215,6 +2231,12 @@ Module DJoin.
   Lemma refl n (a : PTm n) : 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.
+  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.
+  Proof. hauto q:on use:REReds.ToEReds. Qed.
+
   Lemma symmetric n (a b : PTm n) : R a b -> R b a.
   Proof. sfirstorder unfold:R. Qed.
 
@@ -2381,6 +2403,45 @@ Module DJoin.
     hauto q:on ctrs:rtc inv:option use:REReds.cong.
   Qed.
 
+  Lemma pair_inj n (a0 a1 b0 b1 : PTm n) :
+    SN (PPair a0 b0) ->
+    SN (PPair a1 b1) ->
+    R (PPair a0 b0) (PPair a1 b1) ->
+    R a0 a1 /\ R b0 b1.
+  Proof.
+    move => sn0 sn1.
+    have [? [? [? ?]]] : SN a0 /\ SN b0 /\ SN a1 /\ SN b1 by sfirstorder use:SN_NoForbid.P_PairInv.
+    have ? : SN (PProj PL (PPair a0 b0)) by hauto lq:on db:sn.
+    have ? : SN (PProj PR (PPair a0 b0)) by hauto lq:on db:sn.
+    have ? : SN (PProj PL (PPair a1 b1)) by hauto lq:on db:sn.
+    have ? : SN (PProj PR (PPair a1 b1)) by hauto lq:on db:sn.
+    have h0L : RRed.R (PProj PL (PPair a0 b0)) a0 by eauto with red.
+    have h0R : RRed.R (PProj PR (PPair a0 b0)) b0 by eauto with red.
+    have h1L : RRed.R (PProj PL (PPair a1 b1)) a1 by eauto with red.
+    have h1R : RRed.R (PProj PR (PPair a1 b1)) b1 by eauto with red.
+    move => h2.
+    move /ProjCong in h2.
+    have h2L := h2 PL.
+    have {h2}h2R := h2 PR.
+    move /FromRRed1 in h0L.
+    move /FromRRed0 in h1L.
+    move /FromRRed1 in h0R.
+    move /FromRRed0 in h1R.
+    split; eauto using transitive.
+  Qed.
+
+  Lemma ejoin_pair_inj n (a0 a1 b0 b1 : PTm n) :
+    nf a0 -> nf b0 -> nf a1 -> nf b1 ->
+    EJoin.R (PPair a0 b0) (PPair a1 b1) ->
+    EJoin.R a0 a1 /\ EJoin.R b0 b1.
+  Proof.
+    move => h0 h1 h2 h3 /FromEJoin.
+    have [? ?] : SN (PPair a0 b0) /\ SN (PPair a1 b1) by sauto lqb:on rew:off use:ne_nf_embed.
+    move => ?.
+    have : R a0 a1 /\ R b0 b1 by hauto l:on use:pair_inj.
+    hauto l:on use:ToEJoin.
+  Qed.
+
   Lemma abs_inj n (a0 a1 : PTm (S n)) :
     SN a0 -> SN a1 ->
     R (PAbs a0) (PAbs a1) ->

From 49a254fbef3a6db0fa46945ba0884775659a6ac8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 19:51:08 -0500
Subject: [PATCH 108/210] Finish the pair pair case

---
 theories/algorithmic.v | 37 ++++++++++++++++++++++++++++++++-----
 1 file changed, 32 insertions(+), 5 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 24ce2b9..f3de8e6 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -945,12 +945,39 @@ Proof.
       sfirstorder.
     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
+      have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1)
+        by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
       apply : CE_HRed; eauto using rtc_refl.
-      admit.
-      (* suff : exists l, l < n /\ algo_metric l a0 b0 /\ algo_metric l a1 b1. *)
-      (* move => [l [hl [hal0 hal1]]]. *)
-      (* apply CE_PairPair. eapply ih; eauto. *)
-      (* by eapply ih; eauto. *)
+      move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0.
+      move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1.
+      move /Sub_Bind_InvR :  (hSu0).
+      move => [i][A2][B2]hE.
+      have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2
+        by eauto using Su_Transitive, Su_Eq.
+      have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2
+        by eauto using Su_Transitive, Su_Eq.
+      have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1.
+      have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1.
+      have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
+      have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
+      move /algo_metric_pair : h sn0 sn1; repeat move/[apply].
+      move => [j][hj][ih0 ih1].
+      have haE : a0 ⇔ a1 by hauto l:on.
+      apply CE_PairPair => //.
+      have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2
+        by hauto l:on use:coqeq_sound_mutual.
+      have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2.
+      apply : T_Conv; eauto.
+      move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2.
+      eapply ih; cycle -1; eauto.
+      apply : T_Conv; eauto.
+      apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2).
+      move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2.
+      move : hE haE. clear.
+      move => h.
+      eapply regularity in h.
+      move : h => [_ [hB _]].
+      eauto using bind_inst.
     + case : b => //=.
       * hauto lq:on use:T_AbsBind_Imp.
       * hauto lq:on rew:off use:T_PairBind_Imp.

From d24991e99425d50619856c08e9730a1cf6c484d5 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 20:41:57 -0500
Subject: [PATCH 109/210] Finish most of the neu abs case

---
 theories/algorithmic.v | 43 +++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 42 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index f3de8e6..f78e1f6 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -886,6 +886,25 @@ Lemma T_Univ_Raise n Γ (a : PTm n) i j :
   Γ ⊢ a ∈ PUniv j.
 Proof. hauto lq:on rew:off use:T_Conv, Su_Univ, wff_mutual. Qed.
 
+Lemma Abs_Pi_Inv n Γ (a : PTm (S n)) A B :
+  Γ ⊢ PAbs a ∈ PBind PPi A B ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B.
+Proof.
+  move => h.
+  have [i hi] : exists i, Γ ⊢ PBind PPi A B ∈ PUniv i by hauto use:regularity.
+  have  [{}i {}hi] : exists i, Γ ⊢ A ∈ PUniv i by hauto use:Bind_Inv.
+  apply : subject_reduction; last apply RRed.AppAbs'.
+  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.
+  by eauto using Wff_Cons'.
+  rewrite -/ren_PTm.
+  by asimpl.
+  rewrite -/ren_PTm.
+  by asimpl.
+Qed.
+
 Lemma coqeq_complete n k (a b : PTm n) :
   algo_metric k a b ->
   (forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) /\
@@ -1023,7 +1042,29 @@ Proof.
     move => + b.
     case : b => //=.
     (* NeuAbs *)
-    + admit.
+    + move => a u halg _ nea. split => // Γ A hu /[dup] ha.
+      move /Abs_Inv => [A0][B0][_]hSu.
+      move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE.
+      have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
+      have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
+      have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i
+        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 {}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.
+      move /Abs_Pi_Inv in ha.
+      move /algo_metric_neu_abs : halg.
+      move => [j0][hj0]halg.
+      apply : CE_HRed; eauto using rtc_refl.
+      eapply CE_NeuAbs => //=.
+      eapply ih; eauto.
     (* NeuPair *)
     + admit.
     (* NeuBind: Impossible *)

From bdba6f50e52aef0d16451bf9922ad5be12a67fa3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 22:43:56 -0500
Subject: [PATCH 110/210] Finish the soundness proof completely

---
 theories/algorithmic.v | 92 +++++++++++++++++++++++++++++++++++++++---
 theories/common.v      |  4 ++
 theories/fp_red.v      | 28 +++++++++++++
 3 files changed, 119 insertions(+), 5 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index f78e1f6..13b94dd 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -686,6 +686,10 @@ Lemma lored_hne_preservation n (a b : PTm n) :
     LoRed.R a b -> ishne a -> ishne b.
 Proof.  induction 1; sfirstorder. Qed.
 
+Lemma loreds_hne_preservation n (a b : PTm n) :
+  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 :
     nsteps LoRed.R k (PBind p a0 b0) C ->
     exists i j a1 b1,
@@ -771,6 +775,20 @@ Proof.
   lia.
 Qed.
 
+Lemma lored_nsteps_app_cong k n (a0 a1 b : PTm n) :
+  nsteps LoRed.R k a0 a1 ->
+  ishne a0 ->
+  nsteps LoRed.R k (PApp a0 b) (PApp a1 b).
+  move => h. move : b.
+  elim : k a0 a1 / h.
+  - sauto.
+  - move => m a0 a1 a2 h0 h1 ih.
+    move => b hneu.
+    apply : nsteps_l; eauto using LoRed.AppCong0.
+    apply LoRed.AppCong0;eauto. move : hneu. clear. case : a0 => //=.
+    apply ih. sfirstorder use:lored_hne_preservation.
+Qed.
+
 Lemma lored_nsteps_pair_inv k n (a0 b0 C : PTm n) :
     nsteps LoRed.R k (PPair a0 b0) C ->
     exists i j a1 b1,
@@ -822,10 +840,38 @@ 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) :
+  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) :
+  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 :
   algo_metric k u (PAbs a0) ->
+  ishne u ->
   exists j, j < k /\ algo_metric j (PApp (ren_PTm shift u) (VarPTm var_zero)) a0.
-Admitted.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 neu.
+  have neva : ne va by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
+  move /lored_nsteps_abs_inv : h1 => [a1][h01]?. subst.
+  exists (k - 1). simpl in *. split. lia.
+  have {}h0 : nsteps LoRed.R i (ren_PTm shift u) (ren_PTm shift va)
+    by eauto using lored_nsteps_renaming.
+  have {}h0 : nsteps LoRed.R i (PApp (ren_PTm shift u) (VarPTm var_zero)) (PApp (ren_PTm shift va) (VarPTm var_zero)).
+  apply lored_nsteps_app_cong => //=.
+  scongruence use:ishne_ren.
+  do 4 eexists. repeat split; eauto.
+  hauto b:on use:ne_nf_ren.
+  have h : EJoin.R (PAbs (PApp (ren_PTm shift va) (VarPTm var_zero))) (PAbs a1) by hauto q:on ctrs:rtc,ERed.R.
+  apply DJoin.ejoin_abs_inj; eauto.
+  hauto b:on drew:off use:ne_nf_ren.
+  simpl in *.  rewrite size_PTm_ren. lia.
+Qed.
 
 Lemma algo_metric_neu_pair n k (a0 b0 : PTm n) u :
   algo_metric k u (PPair a0 b0) ->
@@ -905,6 +951,24 @@ Proof.
   by asimpl.
 Qed.
 
+Lemma Pair_Sig_Inv n Γ (a b : PTm n) A B :
+  Γ ⊢ PPair a b ∈ PBind PSig A B ->
+  Γ ⊢ a ∈ A /\ Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B.
+Proof.
+  move => /[dup] h0 h1.
+  have [i hr] : exists i, Γ ⊢ PBind PSig A B ∈ PUniv i by sfirstorder use:regularity.
+  move /T_Proj1 in h0.
+  move /T_Proj2 in h1.
+  split.
+  hauto lq:on use:subject_reduction ctrs:RRed.R.
+  have hE : Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A by
+    hauto lq:on use:RRed_Eq ctrs:RRed.R.
+  apply : T_Conv.
+  move /subject_reduction : h1. apply.
+  apply RRed.ProjPair.
+  apply : bind_inst; eauto.
+Qed.
+
 Lemma coqeq_complete n k (a b : PTm n) :
   algo_metric k a b ->
   (forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) /\
@@ -1060,13 +1124,32 @@ Proof.
       by apply T_Var.
       rewrite -/ren_PTm. by asimpl.
       move /Abs_Pi_Inv in ha.
-      move /algo_metric_neu_abs : halg.
+      move /algo_metric_neu_abs /(_ nea) : halg.
       move => [j0][hj0]halg.
       apply : CE_HRed; eauto using rtc_refl.
       eapply CE_NeuAbs => //=.
       eapply ih; eauto.
     (* NeuPair *)
-    + admit.
+    + move => a0 b0 u halg _ neu.
+      split => // Γ A hu /[dup] wt.
+      move /Pair_Inv => [A0][B0][ha0][hb0]hU.
+      move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE.
+      have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on.
+      have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+      move /Pair_Sig_Inv : wt => [{}ha0 {}hb0].
+      have /T_Proj1 huL := hu.
+      have /T_Proj2 {hu} huR := hu.
+      move /algo_metric_neu_pair : halg => [j [hj [hL hR]]].
+      have heL : PProj PL u ⇔ a0 by hauto l:on.
+      eapply CE_HRed; eauto using rtc_refl.
+      apply CE_NeuPair; eauto.
+      eapply ih; eauto.
+      apply : T_Conv; eauto.
+      have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i
+        by hauto l:on use:regularity.
+      have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by
+        hauto l:on use:coqeq_sound_mutual.
+      hauto l:on use:bind_inst.
     (* NeuBind: Impossible *)
     + move => b p p0 a /algo_metric_join h _ h0. exfalso.
       move : h h0. clear.
@@ -1131,7 +1214,6 @@ Proof.
         apply : Su_Transitive; eauto.
         move /E_Refl in ha0.
         hauto l:on use:Su_Pi_Proj2.
-        move => [:hsub].
         have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
         split.
         apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
@@ -1210,4 +1292,4 @@ Proof.
                sfirstorder use:bind_inst.
       * sfirstorder use:T_Bot_Imp.
     + sfirstorder use:T_Bot_Imp.
-Admitted.
+Qed.
diff --git a/theories/common.v b/theories/common.v
index ac70322..24e708e 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -91,3 +91,7 @@ Proof. case : a => //=. Qed.
 Definition ispair_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
   ispair (ren_PTm ξ a) = ispair a.
 Proof. case : a => //=. Qed.
+
+Definition ishne_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
+  ishne (ren_PTm ξ a) = ishne a.
+Proof. move : m ξ. elim : n / a => //=. Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 407d3f5..2ab26d5 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1900,6 +1900,22 @@ Module LoRed.
     RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
 
+  Lemma AppAbs' n a (b : PTm n) u :
+    u = (subst_PTm (scons b VarPTm) a) ->
+    R (PApp (PAbs a) b)  u.
+  Proof. move => ->. apply AppAbs. Qed.
+
+  Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
+    R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
+  Proof.
+    move => h. move : m ξ.
+    elim : n a b /h.
+
+    move => n a b m ξ /=.
+    apply AppAbs'. by asimpl.
+    all : try qauto ctrs:R use:ne_nf_ren, ishf_ren.
+  Qed.
+
 End LoRed.
 
 Module LoReds.
@@ -2467,6 +2483,18 @@ Module DJoin.
     eauto.
   Qed.
 
+  Lemma ejoin_abs_inj n (a0 a1 : PTm (S n)) :
+    nf a0 -> nf a1 ->
+    EJoin.R (PAbs a0) (PAbs a1) ->
+    EJoin.R a0 a1.
+  Proof.
+    move => h0 h1.
+    have [? [? [? ?]]] : SN a0 /\ SN a1 /\ SN (PAbs a0) /\ SN (PAbs a1) by sauto lqb:on rew:off use:ne_nf_embed.
+    move /FromEJoin.
+    move /abs_inj.
+    hauto l:on use:ToEJoin.
+  Qed.
+
   Lemma standardization n (a b : PTm n) :
     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.

From 21d9a2ce1bdf58c681e312fe46bbd694333ba316 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 23:00:23 -0500
Subject: [PATCH 111/210] Add standardization_lo

---
 theories/algorithmic.v |  3 +++
 theories/fp_red.v      | 14 ++++++++++++++
 2 files changed, 17 insertions(+)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 13b94dd..66bc889 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -875,7 +875,10 @@ Qed.
 
 Lemma algo_metric_neu_pair n k (a0 b0 : PTm n) 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.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 neu.
 Admitted.
 
 Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 2ab26d5..b8bcc41 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2515,6 +2515,20 @@ Module DJoin.
     hauto q:on use:NeEPars.ToEReds unfold:EJoin.R.
   Qed.
 
+  Lemma standardization_lo n (a b : PTm n) :
+    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.
+    move => /[dup] sna + /[dup] snb.
+    move : standardization; repeat move/[apply].
+    move => [va][vb][hva][hvb][nfva][nfvb]hj.
+    move /LoReds.FromSN : sna => [va' [hva' hva'0]].
+    move /LoReds.FromSN : snb => [vb' [hvb' hvb'0]].
+    exists va', vb'.
+    repeat split => //=.
+    have : va = va' /\ vb = vb' by sfirstorder use:red_uniquenf, LoReds.ToRReds.
+    case; congruence.
+  Qed.
 End DJoin.
 
 

From eaf59fc45e9222375bd436e2e3cef8a117af1e13 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 23:25:32 -0500
Subject: [PATCH 112/210] Finish all cases of algorithmic completeness

---
 theories/algorithmic.v | 31 +++++++++++++++++++++++++++++--
 1 file changed, 29 insertions(+), 2 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 66bc889..865ad6b 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -789,6 +789,19 @@ 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) :
+  nsteps LoRed.R k a0 a1 ->
+  ishne a0 ->
+  nsteps LoRed.R k (PProj p a0) (PProj p a1).
+  move => h. move : p.
+  elim : k a0 a1 / h.
+  - sauto.
+  - move => m a0 a1 a2 h0 h1 ih p hneu.
+    apply : nsteps_l; eauto using LoRed.ProjCong.
+    apply LoRed.ProjCong;eauto. move : hneu. clear. case : a0 => //=.
+    apply ih. sfirstorder use:lored_hne_preservation.
+Qed.
+
 Lemma lored_nsteps_pair_inv k n (a0 b0 C : PTm n) :
     nsteps LoRed.R k (PPair a0 b0) C ->
     exists i j a1 b1,
@@ -879,7 +892,20 @@ Lemma algo_metric_neu_pair n k (a0 b0 : PTm n) u :
   exists j, j < k /\ algo_metric j (PProj PL u) a0 /\ algo_metric j (PProj PR u) b0.
 Proof.
   move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 neu.
-Admitted.
+  move /lored_nsteps_pair_inv : h1.
+  move => [i0][j0][a1][b1][hi][hj][?][ha01]hb01. subst.
+  simpl in *.
+  have ? : ishne va by
+    hauto lq:on use:loreds_hne_preservation, @relations.rtc_nsteps.
+  have ? : ne va by sfirstorder use:hne_nf_ne.
+  exists (k - 1). split. lia.
+  move :lored_nsteps_proj_cong (neu) h0; repeat move/[apply].
+  move => h. have hL := h PL. have {h} hR := h PR.
+  suff [? ?] : EJoin.R (PProj PL va) a1  /\ EJoin.R (PProj PR va) b1.
+  rewrite /algo_metric.
+  split; do 4 eexists; repeat split;eauto; sfirstorder b:on solve+:lia.
+  eapply DJoin.ejoin_pair_inj; hauto qb:on ctrs:rtc, ERed.R.
+Qed.
 
 Lemma algo_metric_app n k (a0 b0 a1 b1 : PTm n) :
   algo_metric k (PApp a0 b0) (PApp a1 b1) ->
@@ -1021,6 +1047,7 @@ Proof.
       move => /[dup] [[ha00 ha10]] [].
       move : DJoin.abs_inj; repeat move/[apply].
       move : DJoin.standardization ha00 ha10; repeat move/[apply].
+      (* this is awful *)
       move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]].
       have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va'
         by hauto lq:on use:@relations.rtc_nsteps.
@@ -1142,7 +1169,7 @@ Proof.
       move /Pair_Sig_Inv : wt => [{}ha0 {}hb0].
       have /T_Proj1 huL := hu.
       have /T_Proj2 {hu} huR := hu.
-      move /algo_metric_neu_pair : halg => [j [hj [hL hR]]].
+      move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]].
       have heL : PProj PL u ⇔ a0 by hauto l:on.
       eapply CE_HRed; eauto using rtc_refl.
       apply CE_NeuPair; eauto.

From 8daaae9831e1dc0f0eceb1167a58b388e37700a0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 23:39:56 -0500
Subject: [PATCH 113/210] Minor

---
 theories/algorithmic.v | 25 ++++++++++++++++++++++++-
 1 file changed, 24 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 865ad6b..444f49f 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -998,7 +998,7 @@ Proof.
   apply : bind_inst; eauto.
 Qed.
 
-Lemma coqeq_complete n k (a b : PTm n) :
+Lemma coqeq_complete' n k (a b : PTm n) :
   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).
@@ -1323,3 +1323,26 @@ Proof.
       * sfirstorder use:T_Bot_Imp.
     + sfirstorder use:T_Bot_Imp.
 Qed.
+
+Lemma coqeq_sound : forall n Γ (a b : PTm n) 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) :
+  Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
+Proof.
+  move => h.
+  suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity.
+  eapply fundamental_theorem in h.
+  move /logrel.SemEq_SN_Join : h.
+  move => h.
+  have : exists va vb : PTm n,
+         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.
+  move => [va][vb][hva][hvb][nva][nvb]hj.
+  move /relations.rtc_nsteps : hva => [i hva].
+  move /relations.rtc_nsteps : hvb => [j hvb].
+  exists (i + j + size_PTm va + size_PTm vb).
+  hauto lq:on solve+:lia.
+Qed.

From ef3de3af3d68a1d768226f9ef0f9935fa560de6b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 16 Feb 2025 23:53:14 -0500
Subject: [PATCH 114/210] Add the specification for algorithmic subtyping

---
 theories/algorithmic.v | 36 ++++++++++++++++++++++++++++++++++--
 1 file changed, 34 insertions(+), 2 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 444f49f..96b7cb5 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -226,8 +226,6 @@ Qed.
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
 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 :=
 | CE_AbsAbs a b :
   a ⇔ b ->
@@ -1346,3 +1344,37 @@ Proof.
   exists (i + j + size_PTm va + size_PTm vb).
   hauto lq:on solve+:lia.
 Qed.
+
+
+Reserved Notation "a ≪ b" (at level 70).
+Reserved Notation "a ⋖ b" (at level 70).
+Inductive CoqLEq {n} : PTm n -> PTm n -> Prop :=
+| CLE_UnivCong i j :
+  i <= j ->
+  (* -------------------------- *)
+  PUniv i ⋖ PUniv j
+
+| CLE_PiCong A0 A1 B0 B1 :
+  A1 ≪  A0 ->
+  B0 ≪ B1 ->
+  (* ---------------------------- *)
+  PBind PPi A0 B0 ⋖ PBind PPi A1 B1
+
+| CLE_SigCong A0 A1 B0 B1 :
+  A0 ≪ A1 ->
+  B0 ≪ B1 ->
+  (* ---------------------------- *)
+  PBind PSig A0 B0 ⋖ PBind PSig A1 B1
+
+| CLE_NeuNeu a0 a1 :
+  a0 ∼ a1 ->
+  a0 ⋖ a1
+
+with CoqLEq_R {n} : PTm n -> PTm n -> Prop :=
+| CLE_HRed a a' b b'  :
+  rtc HRed.R a a' ->
+  rtc HRed.R b b' ->
+  a' ⋖ b' ->
+  (* ----------------------- *)
+  a ≪ b
+where "a ≪ b" := (CoqLEq_R a b) and "a ⋖ b" := (CoqLEq a b).

From 067ae89b1dd6e66baf654b0773c5c5162695fb85 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 17 Feb 2025 18:34:48 -0500
Subject: [PATCH 115/210] Finish soundness for subtyping

---
 theories/algorithmic.v | 48 ++++++++++++++++++++++++++++++++++++++++++
 theories/fp_red.v      |  4 ++++
 2 files changed, 52 insertions(+)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 96b7cb5..1576627 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -959,6 +959,16 @@ Lemma T_Univ_Raise n Γ (a : PTm n) 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 :
+  Γ ⊢ PBind p A B ∈ PUniv i ->
+  Γ ⊢ A ∈ PUniv i /\ funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i.
+Proof.
+  move /Bind_Inv.
+  move => [i0][hA][hB]h.
+  move /synsub_to_usub : h => [_ [_ /Sub.univ_inj ? ]].
+  sfirstorder use:T_Univ_Raise.
+Qed.
+
 Lemma Abs_Pi_Inv n Γ (a : PTm (S n)) A B :
   Γ ⊢ PAbs a ∈ PBind PPi A B ->
   funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B.
@@ -1378,3 +1388,41 @@ with CoqLEq_R {n} : PTm n -> PTm n -> Prop :=
   (* ----------------------- *)
   a ≪ b
 where "a ≪ b" := (CoqLEq_R a b) and "a ⋖ b" := (CoqLEq a b).
+
+Scheme coqleq_ind := Induction for CoqLEq Sort Prop
+  with coqleq_r_ind := Induction for CoqLEq_R Sort Prop.
+
+Combined Scheme coqleq_mutual from coqleq_ind, coqleq_r_ind.
+
+Definition salgo_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 /\ ESub.R va vb /\ size_PTm va + size_PTm vb + i + j <= k.
+
+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 ).
+Proof.
+  apply coqleq_mutual.
+  - hauto lq:on use:wff_mutual ctrs:LEq.
+  - move => n 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.
+    apply Su_Transitive with (B := PBind PPi A1 B0).
+    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 /Bind_Univ_Inv => [hA0]hB0 /Bind_Univ_Inv [hA1]hB1.
+    have hlA  : Γ ⊢ A0 ≲ A1 by sfirstorder.
+    have hΓ : ⊢ Γ by sfirstorder use:wff_mutual.
+    apply Su_Transitive with (B := PBind PSig A0 B1).
+    apply : Su_Sig; eauto using E_Refl, Su_Eq.
+    apply : ihB; by eauto using ctx_eq_subst_one.
+    apply : Su_Sig; eauto using E_Refl, Su_Eq.
+  - sauto lq:on use:coqeq_sound_mutual, Su_Eq.
+  - move => n 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.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index b8bcc41..0e597e2 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2780,3 +2780,7 @@ Module Sub.
     hauto ctrs:rtc use:REReds.cong', Sub1.substing.
   Qed.
 End Sub.
+
+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.
+End ESub.

From 735c7f20469e8014c1ea6142672bee79e0d6e918 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 17 Feb 2025 21:43:21 -0500
Subject: [PATCH 116/210] Prove some simple soundness cases of subtyping

---
 theories/algorithmic.v | 138 +++++++++++++++++++++++++++++++++++++++++
 theories/fp_red.v      |  14 +++++
 2 files changed, 152 insertions(+)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 1576627..dc7f4b0 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -657,6 +657,18 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
+Lemma T_AbsUniv_Imp' n Γ (a : PTm (S n)) 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  :
+  Γ ⊢ 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) :
   Γ ⊢ PAbs a ∈ U ->
   Γ ⊢ PBind p A0 B0 ∈ U ->
@@ -1397,6 +1409,24 @@ Combined Scheme coqleq_mutual from coqleq_ind, coqleq_r_ind.
 Definition salgo_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 /\ ESub.R va vb /\ size_PTm va + size_PTm vb + i + j <= k.
 
+Lemma salgo_metric_algo_metric n k (a b : PTm n) :
+  ishne a \/ ishne b ->
+  salgo_metric k a b ->
+  algo_metric k a b.
+Proof.
+  move => h.
+  move => [i][j][va][vb][hva][hvb][nva][nvb][hS]sz.
+  rewrite/ESub.R in hS.
+  move : hS => [va'][vb'][h0][h1]h2.
+  suff : va' = vb' by sauto lq:on.
+  have {}hva : rtc RERed.R a va by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds.
+  have {}hvb : rtc RERed.R b vb by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds.
+  apply REReds.FromEReds in h0, h1.
+  have : ishne va' \/ ishne vb' by
+    hauto lq:on rew:off use:@relations.rtc_transitive, REReds.hne_preservation.
+  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 ).
@@ -1426,3 +1456,111 @@ Proof.
     suff : Γ ⊢ a' ≲ b' by eauto using Su_Transitive.
     eauto using HReds.preservation.
 Qed.
+
+Lemma salgo_metric_case n k (a b : PTm n) :
+  salgo_metric k a b ->
+  (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ salgo_metric k' a' b /\ k' < k.
+Proof.
+  move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6.
+  case : a h0 => //=; try firstorder.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.AppAbs unfold:algo_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder qb:on use:ne_hne.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+Qed.
+
+Lemma CLE_HRedL n (a a' b : PTm n)  :
+  HRed.R a a' ->
+  a' ≪ b ->
+  a ≪ b.
+Proof.
+  hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
+Qed.
+
+Lemma CLE_HRedR n (a a' b : PTm n)  :
+  HRed.R a a' ->
+  b ≪ a' ->
+  b ≪ a.
+Proof.
+  hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
+Qed.
+
+
+Lemma algo_metric_caseR n k (a b : PTm n) :
+  salgo_metric k a b ->
+  (ishf b \/ ishne b) \/ exists k' b', HRed.R b b' /\ salgo_metric k' a b' /\ k' < k.
+Proof.
+  move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6.
+  case : b h1 => //=; try by firstorder.
+  - inversion 1 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto q:on use:HRed.AppAbs unfold:salgo_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:salgo_metric solve+:lia.
+    + sfirstorder qb:on use:ne_hne.
+  - inversion 1 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+Qed.
+
+Lemma salgo_metric_sub n k (a b : PTm n) :
+  salgo_metric k a b ->
+  Sub.R a b.
+Proof.
+  rewrite /algo_metric.
+  move => [i][j][va][vb][h0][h1][h2][h3][[va' [vb' [hva [hvb hS]]]]]h5.
+  have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
+  have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
+  apply REReds.FromEReds in hva,hvb.
+  apply LoReds.ToRReds in h0,h1.
+  apply REReds.FromRReds in h0,h1.
+  rewrite /Sub.R. exists va', vb'. sfirstorder use:@relations.rtc_transitive.
+Qed.
+
+Lemma coqleq_complete' n k (a b : PTm n) :
+  salgo_metric k a b -> (forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> a ≪ b).
+Proof.
+  move : k n a b.
+  elim /Wf_nat.lt_wf_ind.
+  move => n ih.
+  move => k a b /[dup] h /salgo_metric_case.
+  (* Cases where a and b can take steps *)
+  case; cycle 1.
+  move : k 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.
+  (* Cases where neither a nor b can take steps *)
+  case => fb; case => fa.
+  - case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+      * move => p0 A0 B0 _ p1 A1 B1 _.
+        move => h.
+        have ? : p1 = p0 by
+          hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj.
+        subst.
+        case : p0 h => //=; admit.
+      * hauto lq:on use:salgo_metric_sub, Sub.bind_univ_noconf.
+    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+      * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
+      * move => *. econstructor; eauto using rtc_refl.
+        hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong.
+  (* Both cases are impossible *)
+  - admit.
+  - admit.
+  - move => Γ i ha hb.
+    econstructor; eauto using rtc_refl.
+    apply CLE_NeuNeu. move {ih}.
+    have {}h : algo_metric n a b by
+      hauto lq:on use:salgo_metric_algo_metric.
+    eapply coqeq_complete'; eauto.
+Admitted.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 0e597e2..3a40f2e 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2620,6 +2620,10 @@ Module Sub1.
     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) :
+    ishne a \/ ishne b -> R a b -> a = b.
+  Proof. hauto q:on inv:R. Qed.
+
 End Sub1.
 
 Module Sub.
@@ -2628,6 +2632,16 @@ Module Sub.
   Lemma refl n (a : PTm n) : R a a.
   Proof. sfirstorder use:@rtc_refl unfold:R. Qed.
 
+  Lemma ToJoin n (a b : PTm n) :
+    ishne a \/ ishne b ->
+    R a b ->
+    DJoin.R a b.
+  Proof.
+    move => h [c][d][h0][h1]h2.
+    have : ishne c \/ ishne d by hauto q:on use:REReds.hne_preservation.
+    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.
   Proof.
     rewrite /R.

From 9c5eb31edfdba5e289cf8ac5a7e1e609b9f3728f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 17 Feb 2025 22:50:25 -0500
Subject: [PATCH 117/210] Finish all cases of subtyping

---
 theories/algorithmic.v | 61 +++++++++++++++++++++++++++++++++++++++---
 theories/fp_red.v      | 32 ++++++++++++++++++++++
 2 files changed, 89 insertions(+), 4 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index dc7f4b0..2ea3f78 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -946,6 +946,7 @@ Proof.
     repeat split => //=; sfirstorder b:on use:ne_nf.
 Qed.
 
+
 Lemma algo_metric_bind n k p0 (A0 : PTm n) 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.
@@ -1526,6 +1527,32 @@ 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  :
+  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.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
+  move : lored_nsteps_bind_inv h0 => /[apply].
+  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
+  move : lored_nsteps_bind_inv h1 => /[apply].
+  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+  move /ESub.pi_inj : h4 => [? ?].
+  hauto qb:on solve+:lia.
+Qed.
+
+Lemma salgo_metric_sig n k (A0 : PTm n) 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.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
+  move : lored_nsteps_bind_inv h0 => /[apply].
+  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
+  move : lored_nsteps_bind_inv h1 => /[apply].
+  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+  move /ESub.sig_inj : h4 => [? ?].
+  hauto qb:on solve+:lia.
+Qed.
+
 Lemma coqleq_complete' n k (a b : PTm n) :
   salgo_metric k a b -> (forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> a ≪ b).
 Proof.
@@ -1548,19 +1575,45 @@ Proof.
         have ? : p1 = p0 by
           hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj.
         subst.
-        case : p0 h => //=; admit.
+        case : p0 h => //=.
+        ** move /salgo_metric_pi.
+           move => [j [hj [hA hB]]] Γ i.
+           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
+           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
+             by hauto l:on.
+           eauto using ctx_eq_subst_one.
+        ** move /salgo_metric_sig.
+           move => [j [hj [hA hB]]] Γ i.
+           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
+           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
+             by hauto l:on.
+           eauto using ctx_eq_subst_one.
       * hauto lq:on use:salgo_metric_sub, Sub.bind_univ_noconf.
     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
       * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
       * move => *. econstructor; eauto using rtc_refl.
         hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong.
   (* Both cases are impossible *)
-  - admit.
-  - admit.
+  - have {}h : DJoin.R a b by
+       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join.
+    case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+    + hauto lq:on use:DJoin.hne_bind_noconf.
+    + hauto lq:on use:DJoin.hne_univ_noconf.
+  - have {}h : DJoin.R b a by
+      hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric.
+    case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+    + hauto lq:on use:DJoin.hne_bind_noconf.
+    + hauto lq:on use:DJoin.hne_univ_noconf.
   - move => Γ i ha hb.
     econstructor; eauto using rtc_refl.
     apply CLE_NeuNeu. move {ih}.
     have {}h : algo_metric n a b by
       hauto lq:on use:salgo_metric_algo_metric.
     eapply coqeq_complete'; eauto.
-Admitted.
+Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3a40f2e..d1702fe 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2337,6 +2337,17 @@ Module DJoin.
     case : c => //=.
   Qed.
 
+  Lemma hne_bind_noconf n (a b : PTm n) :
+    R a b -> ishne a -> isbind b -> False.
+  Proof.
+    move => [c [h0 h1]] h2 h3.
+    have {h0 h1 h2 h3} : ishne c /\ isbind c by
+      hauto l:on use:REReds.hne_preservation,
+          REReds.bind_preservation.
+    move => [].
+    case : c => //=.
+  Qed.
+
   Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
     DJoin.R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ DJoin.R A0 A1 /\ DJoin.R B0 B1.
@@ -2797,4 +2808,25 @@ End Sub.
 
 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.
+
+  Lemma pi_inj n (A0 A1 : PTm n) B0 B1 :
+    R (PBind PPi A0 B0) (PBind PPi A1 B1) ->
+    R A1 A0 /\ R B0 B1.
+  Proof.
+    move => [u0 [u1 [h0 [h1 h2]]]].
+    move /EReds.bind_inv : h0 => [A2][B2][?][h3]h4. subst.
+    move /EReds.bind_inv : h1 => [A3][B3][?][h5]h6. subst.
+    sauto lq:on rew:off inv:Sub1.R.
+  Qed.
+
+  Lemma sig_inj n (A0 A1 : PTm n) B0 B1 :
+    R (PBind PSig A0 B0) (PBind PSig A1 B1) ->
+    R A0 A1 /\ R B0 B1.
+  Proof.
+    move => [u0 [u1 [h0 [h1 h2]]]].
+    move /EReds.bind_inv : h0 => [A2][B2][?][h3]h4. subst.
+    move /EReds.bind_inv : h1 => [A3][B3][?][h5]h6. subst.
+    sauto lq:on rew:off inv:Sub1.R.
+  Qed.
+
 End ESub.

From d48d9db1b74cbd044ac89d36ec29904c2d755fd0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 17 Feb 2025 23:31:12 -0500
Subject: [PATCH 118/210] Finish the conversion proof completely

---
 theories/algorithmic.v | 29 ++++++++++++++
 theories/fp_red.v      | 87 ++++++++++++++++++++++++++++++------------
 2 files changed, 91 insertions(+), 25 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 2ea3f78..6fe5fad 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1617,3 +1617,32 @@ Proof.
       hauto lq:on use:salgo_metric_algo_metric.
     eapply coqeq_complete'; eauto.
 Qed.
+
+Lemma coqleq_complete n Γ (a b : PTm n) :
+  Γ ⊢ a ≲ b -> a ≪ b.
+Proof.
+  move => h.
+  suff : exists k, salgo_metric k a b by hauto lq:on use:coqleq_complete', regularity.
+  eapply fundamental_theorem in h.
+  move /logrel.SemLEq_SN_Sub : h.
+  move => h.
+  have : exists va vb : PTm n,
+         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.
+  move => [va][vb][hva][hvb][nva][nvb]hj.
+  move /relations.rtc_nsteps : hva => [i hva].
+  move /relations.rtc_nsteps : hvb => [j hvb].
+  exists (i + j + size_PTm va + size_PTm vb).
+  hauto lq:on solve+:lia.
+Qed.
+
+Lemma coqleq_sound : forall n Γ (a b : PTm n) i j,
+    Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv j -> a ≪ b -> Γ ⊢ a ≲ b.
+Proof.
+  move => n Γ 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.
+  sfirstorder use:coqleq_sound_mutual.
+Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index d1702fe..e89c191 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2637,6 +2637,31 @@ Module Sub1.
 
 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.
+
+  Lemma pi_inj n (A0 A1 : PTm n) B0 B1 :
+    R (PBind PPi A0 B0) (PBind PPi A1 B1) ->
+    R A1 A0 /\ R B0 B1.
+  Proof.
+    move => [u0 [u1 [h0 [h1 h2]]]].
+    move /EReds.bind_inv : h0 => [A2][B2][?][h3]h4. subst.
+    move /EReds.bind_inv : h1 => [A3][B3][?][h5]h6. subst.
+    sauto lq:on rew:off inv:Sub1.R.
+  Qed.
+
+  Lemma sig_inj n (A0 A1 : PTm n) B0 B1 :
+    R (PBind PSig A0 B0) (PBind PSig A1 B1) ->
+    R A0 A1 /\ R B0 B1.
+  Proof.
+    move => [u0 [u1 [h0 [h1 h2]]]].
+    move /EReds.bind_inv : h0 => [A2][B2][?][h3]h4. subst.
+    move /EReds.bind_inv : h1 => [A3][B3][?][h5]h6. subst.
+    sauto lq:on rew:off inv:Sub1.R.
+  Qed.
+
+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.
 
@@ -2804,29 +2829,41 @@ Module Sub.
     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.
+  Proof. hauto q:on use:REReds.ToEReds. Qed.
+
+  Lemma standardization n (a b : PTm n) :
+    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.
+    move => h0 h1 hS.
+    have : exists v, rtc RRed.R a v  /\ nf v by sfirstorder use:LoReds.FromSN, LoReds.ToRReds.
+    move => [v [hv2 hv3]].
+    have : exists v, rtc RRed.R b v  /\ nf v by sfirstorder use:LoReds.FromSN, LoReds.ToRReds.
+    move => [v' [hv2' hv3']].
+    move : (hv2) (hv2') => *.
+    apply DJoin.FromRReds in hv2, hv2'.
+    move/DJoin.symmetric in hv2'.
+    apply FromJoin in hv2, hv2'.
+    have hv : R v v' by eauto using transitive.
+    have {}hv : ESub.R v v' by hauto l:on use:ToESub.
+    hauto lq:on.
+  Qed.
+
+  Lemma standardization_lo n (a b : PTm n) :
+    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.
+    move => /[dup] sna + /[dup] snb.
+    move : standardization; repeat move/[apply].
+    move => [va][vb][hva][hvb][nfva][nfvb]hj.
+    move /LoReds.FromSN : sna => [va' [hva' hva'0]].
+    move /LoReds.FromSN : snb => [vb' [hvb' hvb'0]].
+    exists va', vb'.
+    repeat split => //=.
+    have : va = va' /\ vb = vb' by sfirstorder use:red_uniquenf, LoReds.ToRReds.
+    case; congruence.
+  Qed.
+
 End Sub.
-
-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.
-
-  Lemma pi_inj n (A0 A1 : PTm n) B0 B1 :
-    R (PBind PPi A0 B0) (PBind PPi A1 B1) ->
-    R A1 A0 /\ R B0 B1.
-  Proof.
-    move => [u0 [u1 [h0 [h1 h2]]]].
-    move /EReds.bind_inv : h0 => [A2][B2][?][h3]h4. subst.
-    move /EReds.bind_inv : h1 => [A3][B3][?][h5]h6. subst.
-    sauto lq:on rew:off inv:Sub1.R.
-  Qed.
-
-  Lemma sig_inj n (A0 A1 : PTm n) B0 B1 :
-    R (PBind PSig A0 B0) (PBind PSig A1 B1) ->
-    R A0 A1 /\ R B0 B1.
-  Proof.
-    move => [u0 [u1 [h0 [h1 h2]]]].
-    move /EReds.bind_inv : h0 => [A2][B2][?][h3]h4. subst.
-    move /EReds.bind_inv : h1 => [A3][B3][?][h5]h6. subst.
-    sauto lq:on rew:off inv:Sub1.R.
-  Qed.
-
-End ESub.

From df0b955e4ee41eebfc57332bb5d0851f437edcd3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 19 Feb 2025 16:04:02 -0500
Subject: [PATCH 119/210] Add the stub for Equations but might give up later

---
 theories/executable.v | 20 +++++++++++++++++++
 theories/fp_red.v     | 45 ++++++++++++++++++++++++++++++++++++++++++-
 2 files changed, 64 insertions(+), 1 deletion(-)
 create mode 100644 theories/executable.v

diff --git a/theories/executable.v b/theories/executable.v
new file mode 100644
index 0000000..34a847d
--- /dev/null
+++ b/theories/executable.v
@@ -0,0 +1,20 @@
+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 ssreflect ssrbool.
+
+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.
+
+Search (nat -> nat -> bool).
+
+Equations check_equal {n k} (a b : PTm n) (h : metric k a b) :
+  bool by wf k lt :=
+  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) _ := _;
+  check_equal _ _ _ := _.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index e89c191..3022bea 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1395,7 +1395,50 @@ Module ERed.
     all : qauto ctrs:R.
   Qed.
 
-  (* Need to generalize to injective renaming *)
+  (* 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.
+
+  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).
+  Proof.
+    move => hξ.
+    destruct j. asimpl.
+    move => h. have :  i<> f by hauto lq:on rew:off inv:option.
+    move /hξ. by rewrite /funcomp => ->.
+    done.
+  Qed.
+
+  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.
+  Proof.
+    rewrite /tm_i_free.
+    move => [+ [+ [+ +]]].
+    move : i.
+    elim : n / a => n.
+    - hauto q:on.
+    - move => a iha i m ρ0 ρ1 [] => h [] h' m' ξ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.
+    - 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ξ /=.
+      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_PTm ρ0) (upRen_PTm_PTm ρ1)).
+      apply. hauto l:on.
+    - hauto lq:on rew:off.
+    - hauto lq:on rew:off.
+  Qed.
+
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : 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.

From fe5c16361abc3ebfb09befc36cc37677532d5f96 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 19 Feb 2025 17:40:56 -0500
Subject: [PATCH 120/210] Add definition for algorithmic domain

---
 theories/executable.v | 84 +++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 84 insertions(+)

diff --git a/theories/executable.v b/theories/executable.v
index 34a847d..47b9764 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -5,6 +5,90 @@ Require Import algorithmic.
 From stdpp Require Import relations (rtc(..), nsteps(..)).
 Require Import ssreflect ssrbool.
 
+Inductive algo_dom {n} : PTm n -> PTm n -> Prop :=
+| A_AbsAbs a b :
+  algo_dom 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 (PAbs a) u
+
+| A_NeuAbs a u :
+  ishne u ->
+  algo_dom (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 (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 (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 u (PPair a0 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 (PBind p0 A0 B0) (PBind p1 A1 B1)
+
+| A_VarCong i j :
+  (* -------------------------- *)
+  algo_dom (VarPTm i) (VarPTm j)
+
+| A_ProjCong p0 p1 u0 u1 :
+  ishne u0 ->
+  ishne u1 ->
+  algo_dom u0 u1 ->
+  (* ---------------------  *)
+  algo_dom (PProj p0 u0) (PProj p1 u1)
+
+| A_AppCong u0 u1 a0 a1 :
+  ishne u0 ->
+  ishne u1 ->
+  algo_dom u0 u1 ->
+  algo_dom a0 a1 ->
+  (* ------------------------- *)
+  algo_dom (PApp u0 a0) (PApp u1 a1)
+
+| A_HRedL a a' b  :
+  HRed.R a a' ->
+  algo_dom a' b ->
+  (* ----------------------- *)
+  algo_dom a b
+
+| A_HRedR a b b'  :
+  ishne a \/ ishf a ->
+  HRed.R b b' ->
+  algo_dom a b' ->
+  (* ----------------------- *)
+  algo_dom a b.
+
+Search (Bool.reflect (is_true _) _).
+Check idP.
+
+
 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.

From 0e0d9b20e5aa8d7ccc74ece9dde0e85efa9aaaf1 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 19 Feb 2025 18:03:32 -0500
Subject: [PATCH 121/210] Try making the cases mutually recursive?

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

diff --git a/theories/executable.v b/theories/executable.v
index 47b9764..f773fa6 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -85,20 +85,61 @@ Inductive algo_dom {n} : PTm n -> PTm n -> Prop :=
   (* ----------------------- *)
   algo_dom a b.
 
+
+Definition fin_eq {n} (i j : fin n) : bool.
+Proof.
+  induction n.
+  - by exfalso.
+  - refine (match i , j with
+            | None, None => true
+            | Some i, Some j => IHn i j
+            | _, _ => false
+            end).
+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.
+
+
+Equations check_equal {n} (a b : PTm n) (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 (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.
+
 Search (Bool.reflect (is_true _) _).
 Check idP.
 
-
 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.
 
 Search (nat -> nat -> bool).
-
-Equations check_equal {n k} (a b : PTm n) (h : metric k a b) :
-  bool by wf k lt :=
-  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) _ := _;
-  check_equal _ _ _ := _.

From fd0b48073d12cc9b97f401ac95a917d6bfa34267 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 21 Feb 2025 13:23:38 -0500
Subject: [PATCH 122/210] Add nat type definition

---
 syntax.sig                   |   6 +-
 theories/Autosubst2/syntax.v | 165 ++++++++++++++++++++++++++++++++++-
 theories/common.v            |  40 ++++-----
 3 files changed, 184 insertions(+), 27 deletions(-)

diff --git a/syntax.sig b/syntax.sig
index 6b7e4df..a50911e 100644
--- a/syntax.sig
+++ b/syntax.sig
@@ -16,4 +16,8 @@ PPair : PTm -> PTm -> PTm
 PProj : PTag -> PTm -> PTm
 PBind : BTag -> PTm -> (bind PTm in PTm) -> PTm
 PUniv : nat -> PTm
-PBot : PTm
\ No newline at end of file
+PBot : PTm
+PNat : PTm
+PZero : PTm
+PSuc : PTm -> PTm
+PInd : (bind PTm in PTm) -> PTm -> PTm -> (bind PTm,PTm in PTm) -> PTm
\ No newline at end of file
diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index ff9ec18..aae3e02 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -41,7 +41,13 @@ Inductive PTm (n_PTm : nat) : Type :=
   | 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.
+  | 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.
 
 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.
@@ -95,6 +101,37 @@ Proof.
 exact (eq_refl).
 Qed.
 
+Lemma congr_PNat {m_PTm : nat} : PNat m_PTm = PNat m_PTm.
+Proof.
+exact (eq_refl).
+Qed.
+
+Lemma congr_PZero {m_PTm : nat} : PZero m_PTm = PZero m_PTm.
+Proof.
+exact (eq_refl).
+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.
+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)).
+Qed.
+
 Lemma upRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n) :
   fin (S m) -> fin (S n).
 Proof.
@@ -119,6 +156,13 @@ Fixpoint ren_PTm {m_PTm : nat} {n_PTm : nat}
       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)
+        (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) :
@@ -150,6 +194,13 @@ Fixpoint subst_PTm {m_PTm : nat} {n_PTm : nat}
         (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)
+        (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)
@@ -193,6 +244,15 @@ Fixpoint idSubst_PTm {m_PTm : nat} (sigma_PTm : fin m_PTm -> PTm m_PTm)
         (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 =>
+      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)
+        (idSubst_PTm (up_PTm_PTm (up_PTm_PTm sigma_PTm))
+           (upId_PTm_PTm _ (upId_PTm_PTm _ Eq_PTm)) s3)
   end.
 
 Lemma upExtRen_PTm_PTm {m : nat} {n : nat} (xi : fin m -> fin n)
@@ -239,6 +299,18 @@ ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
            (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 =>
+      congr_PInd
+        (extRen_PTm (upRen_PTm_PTm xi_PTm) (upRen_PTm_PTm zeta_PTm)
+           (upExtRen_PTm_PTm _ _ Eq_PTm) s0)
+        (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
+        (extRen_PTm xi_PTm zeta_PTm Eq_PTm s2)
+        (extRen_PTm (upRen_PTm_PTm (upRen_PTm_PTm xi_PTm))
+           (upRen_PTm_PTm (upRen_PTm_PTm zeta_PTm))
+           (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)
@@ -286,6 +358,18 @@ subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
            (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 =>
+      congr_PInd
+        (ext_PTm (up_PTm_PTm sigma_PTm) (up_PTm_PTm tau_PTm)
+           (upExt_PTm_PTm _ _ Eq_PTm) s0)
+        (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
+        (ext_PTm sigma_PTm tau_PTm Eq_PTm s2)
+        (ext_PTm (up_PTm_PTm (up_PTm_PTm sigma_PTm))
+           (up_PTm_PTm (up_PTm_PTm tau_PTm))
+           (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)
@@ -334,6 +418,20 @@ Fixpoint compRenRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
            (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 =>
+      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)
+        (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
+        (compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s2)
+        (compRenRen_PTm (upRen_PTm_PTm (upRen_PTm_PTm xi_PTm))
+           (upRen_PTm_PTm (upRen_PTm_PTm zeta_PTm))
+           (upRen_PTm_PTm (upRen_PTm_PTm rho_PTm))
+           (up_ren_ren _ _ _ (up_ren_ren _ _ _ Eq_PTm)) s3)
   end.
 
 Lemma up_ren_subst_PTm_PTm {k : nat} {l : nat} {m_PTm : nat}
@@ -391,6 +489,21 @@ Fixpoint compRenSubst_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
            (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 =>
+      congr_PSuc (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s0)
+  | 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)
+        (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
+        (compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s2)
+        (compRenSubst_PTm (upRen_PTm_PTm (upRen_PTm_PTm xi_PTm))
+           (up_PTm_PTm (up_PTm_PTm tau_PTm))
+           (up_PTm_PTm (up_PTm_PTm theta_PTm))
+           (up_ren_subst_PTm_PTm _ _ _ (up_ren_subst_PTm_PTm _ _ _ Eq_PTm))
+           s3)
   end.
 
 Lemma up_subst_ren_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
@@ -468,6 +581,21 @@ ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
            (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 =>
+      congr_PSuc (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s0)
+  | 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)
+        (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
+        (compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s2)
+        (compSubstRen_PTm (up_PTm_PTm (up_PTm_PTm sigma_PTm))
+           (upRen_PTm_PTm (upRen_PTm_PTm zeta_PTm))
+           (up_PTm_PTm (up_PTm_PTm theta_PTm))
+           (up_subst_ren_PTm_PTm _ _ _ (up_subst_ren_PTm_PTm _ _ _ Eq_PTm))
+           s3)
   end.
 
 Lemma up_subst_subst_PTm_PTm {k : nat} {l_PTm : nat} {m_PTm : nat}
@@ -547,6 +675,21 @@ subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
            (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 =>
+      congr_PSuc (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s0)
+  | 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)
+        (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
+        (compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s2)
+        (compSubstSubst_PTm (up_PTm_PTm (up_PTm_PTm sigma_PTm))
+           (up_PTm_PTm (up_PTm_PTm tau_PTm))
+           (up_PTm_PTm (up_PTm_PTm theta_PTm))
+           (up_subst_subst_PTm_PTm _ _ _
+              (up_subst_subst_PTm_PTm _ _ _ Eq_PTm)) s3)
   end.
 
 Lemma renRen_PTm {k_PTm : nat} {l_PTm : nat} {m_PTm : nat}
@@ -665,6 +808,18 @@ Fixpoint rinst_inst_PTm {m_PTm : nat} {n_PTm : nat}
            (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 =>
+      congr_PInd
+        (rinst_inst_PTm (upRen_PTm_PTm xi_PTm) (up_PTm_PTm sigma_PTm)
+           (rinstInst_up_PTm_PTm _ _ Eq_PTm) s0)
+        (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
+        (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s2)
+        (rinst_inst_PTm (upRen_PTm_PTm (upRen_PTm_PTm xi_PTm))
+           (up_PTm_PTm (up_PTm_PTm sigma_PTm))
+           (rinstInst_up_PTm_PTm _ _ (rinstInst_up_PTm_PTm _ _ Eq_PTm)) s3)
   end.
 
 Lemma rinstInst'_PTm {m_PTm : nat} {n_PTm : nat}
@@ -871,6 +1026,14 @@ Core.
 
 Arguments VarPTm {n_PTm}.
 
+Arguments PInd {n_PTm}.
+
+Arguments PSuc {n_PTm}.
+
+Arguments PZero {n_PTm}.
+
+Arguments PNat {n_PTm}.
+
 Arguments PBot {n_PTm}.
 
 Arguments PUniv {n_PTm}.
diff --git a/theories/common.v b/theories/common.v
index 24e708e..3c4f0a5 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -7,16 +7,6 @@ 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).
 
-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 q:on depth:2))
-  | _ => solve_anti_ren ()
-  end.
-
-Local Ltac solve_anti_ren := ltac2:(Control.enter solve_anti_ren).
-
 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.
@@ -24,26 +14,22 @@ Proof.
   sblast inv:option.
 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))
+  | _ => 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) ->
   ren_PTm ξ a = ren_PTm ξ b ->
   a = b.
 Proof.
-  move : m ξ b.
-  elim : n / a => //; try solve_anti_ren.
-
-  move => n a iha m ξ []//=.
-  move => u hξ [h].
-  apply iha in h. by subst.
-  destruct i, j=>//=.
-  hauto l:on.
-
-  move => n p A ihA B ihB m ξ []//=.
-  move => b A0 B0  hξ [?]. subst.
-  move => ?. have ? : A0 = A by firstorder. subst.
-  move => ?. have : B = B0. apply : ihB; eauto.
-  sauto.
-  congruence.
+  move : m ξ b. elim : n / a => //; try solve_anti_ren.
 Qed.
 
 Definition ishf {n} (a : PTm n) :=
@@ -52,6 +38,9 @@ Definition ishf {n} (a : PTm n) :=
   | PAbs _ => true
   | PUniv _ => true
   | PBind _ _ _ => true
+  | PNat => true
+  | PSuc _ => true
+  | PZero => true
   | _ => false
   end.
 
@@ -61,6 +50,7 @@ Fixpoint ishne {n} (a : PTm n) :=
   | PApp a _ => ishne a
   | PProj _ a => ishne a
   | PBot => true
+  | PInd _ n _ _ => ishne n
   | _ => false
   end.
 

From 396bddc8b3bebfd6066960114b645763a08c7f3f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 21 Feb 2025 14:35:34 -0500
Subject: [PATCH 123/210] Finish unmorphing

---
 theories/common.v |  21 ++++++
 theories/fp_red.v | 166 ++++++++++++++++++++++++++++++++++++++--------
 2 files changed, 159 insertions(+), 28 deletions(-)

diff --git a/theories/common.v b/theories/common.v
index 3c4f0a5..9cce0cc 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -32,6 +32,9 @@ Proof.
   move : m ξ b. elim : n / a => //; try solve_anti_ren.
 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) :=
   match a with
   | PPair _ _ => true
@@ -44,6 +47,18 @@ Definition ishf {n} (a : PTm n) :=
   | _ => false
   end.
 
+Definition toHF {n} (a : PTm n) :=
+  match a with
+  | PPair _ _ => H_Pair
+  | PAbs _ => H_Abs
+  | PUniv _ => H_Univ
+  | PBind p _ _ => H_Bind p
+  | PNat => H_Nat
+  | PSuc _ => H_Suc
+  | PZero => H_Zero
+  | _ => H_Bot
+  end.
+
 Fixpoint ishne {n} (a : PTm n) :=
   match a with
   | VarPTm _ => true
@@ -64,6 +79,12 @@ Definition ispair {n} (a : PTm n) :=
   | _ => false
   end.
 
+Definition isnat {n} (a : PTm n) := if a is PNat then true else false.
+
+Definition iszero {n} (a : PTm n) := if a is PZero then true else false.
+
+Definition issuc {n} (a : PTm n) := if a is PSuc _ then true else false.
+
 Definition isabs {n} (a : PTm n) :=
   match a with
   | PAbs _ => true
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 3022bea..7068cb1 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -55,7 +55,32 @@ Module EPar.
     R B0 B1 ->
     R (PBind p A0 B0) (PBind p A1 B1)
   | BotCong :
-    R PBot PBot.
+    R PBot PBot
+  | NatCong :
+    R PNat PNat
+  | IndCong P0 P1 a0 a1 b0 b1 c0 c1 :
+    R P0 P1 ->
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    (* ----------------------- *)
+    R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1)
+  | ZeroCong :
+    R PZero PZero
+  | SucCong a0 a1 :
+    R a0 a1 ->
+    (* ------------ *)
+    R (PSuc a0) (PSuc a1)
+  (* | IndZero P b0 b1  c : *)
+  (*   R b0 b1 -> *)
+  (*   R (PInd P PZero b0 c) b1 *)
+  (* | IndSuc P0 P1 a0 a1 b0 b1 c0 c1 : *)
+  (*   R P0 P1 -> *)
+  (*   R a0 a1 -> *)
+  (*   R b0 b1 -> *)
+  (*   R c0 c1 -> *)
+  (*   (* ----------------------------- *) *)
+  (*   R (PInd P0 (PSuc a0) b0 c0) (subst_PTm (scons (PInd P1 a1 b1 c1) (scons a1 VarPTm)) c1) *).
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -76,9 +101,10 @@ Module EPar.
     move => h. move : m ξ.
     elim : n a b /h.
 
+    all : try qauto ctrs:R.
+
     move => n a0 a1 ha iha m ξ /=.
     eapply AppEta'; eauto. by asimpl.
-    all : qauto ctrs:R.
   Qed.
 
   Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
@@ -125,6 +151,12 @@ Inductive SNe {n} : PTm n -> Prop :=
   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 :=
 | N_Pair a b :
   SN a ->
@@ -146,6 +178,13 @@ with SN  {n} : PTm n -> Prop :=
   SN (PBind p A B)
 | N_Univ i :
   SN (PUniv i)
+| N_Nat :
+  SN PNat
+| N_Zero :
+  SN PZero
+| N_Suc a :
+  SN a ->
+  SN (PSuc a)
 with TRedSN {n} : PTm n -> PTm n -> Prop :=
 | N_β a b :
   SN b ->
@@ -162,7 +201,24 @@ with TRedSN {n} : PTm n -> PTm n -> Prop :=
   TRedSN (PProj PR (PPair a b)) b
 | N_ProjCong p a b :
   TRedSN a b ->
-  TRedSN (PProj p a) (PProj p b).
+  TRedSN (PProj p a) (PProj p b)
+| N_IndZero P b c :
+  SN P ->
+  SN b ->
+  SN c ->
+  TRedSN (PInd P PZero b c) b
+| N_IndSuc P a b c :
+  SN P ->
+  SN a ->
+  SN b ->
+  SN c ->
+  TRedSN (PInd P (PSuc a) b c) (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c)
+| N_IndCong P a0 a1 b c :
+  SN P ->
+  SN b ->
+  SN c ->
+  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.
 
@@ -179,7 +235,7 @@ Proof.
   hauto lq:on inv:TRedSN.
 Qed.
 
-Lemma PAbs_imp n a b :
+Lemma PApp_imp n a b :
   @ishf n a ->
   ~~ isabs a ->
   ~ SN (PApp a b).
@@ -190,34 +246,35 @@ Proof.
   hauto lq:on inv:TRedSN.
 Qed.
 
+Lemma PInd_imp n P (a : PTm n) b c :
+  ishf a ->
+  ~~ iszero a ->
+  ~~ issuc a ->
+  ~ SN (PInd P a 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 => //=.
+  hauto lq:on inv:SNe,PTm.
+  hauto lq:on inv:TRedSN.
+Qed.
+
 Lemma PProjAbs_imp n p (a : PTm (S n)) :
   ~ SN (PProj p (PAbs a)).
 Proof.
-  move E : (PProj p (PAbs a)) => u hu.
-  move : p a E.
-  elim : n u / hu=>//=.
-  hauto lq:on inv:SNe.
-  hauto lq:on inv:TRedSN.
+  sfirstorder use:PProj_imp.
 Qed.
 
 Lemma PAppPair_imp n (a b0 b1 : PTm n ) :
   ~ SN (PApp (PPair b0 b1) a).
 Proof.
-  move E : (PApp (PPair b0 b1) a) => u hu.
-  move : a b0 b1 E.
-  elim : n u / hu=>//=.
-  hauto lq:on inv:SNe.
-  hauto lq:on inv:TRedSN.
+  sfirstorder use:PApp_imp.
 Qed.
 
 Lemma PAppBind_imp n p (A : PTm n) B b :
   ~ SN (PApp (PBind p A B) b).
 Proof.
-  move E :(PApp (PBind p A B) b) => u hu.
-  move : p A B b E.
-  elim : n u /hu=> //=.
-  hauto lq:on inv:SNe.
-  hauto lq:on inv:TRedSN.
+  sfirstorder use:PApp_imp.
 Qed.
 
 Lemma PProjBind_imp n p p' (A : PTm n) B :
@@ -246,6 +303,10 @@ Fixpoint ne {n} (a : PTm n) :=
   | 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) :=
   match a with
@@ -257,6 +318,10 @@ with nf {n} (a : PTm n) :=
   | 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.
@@ -289,6 +354,15 @@ Lemma N_β' n a (b : PTm n) u :
   TRedSN (PApp (PAbs a) b) u.
 Proof. move => ->. apply N_β. Qed.
 
+Lemma N_IndSuc' n P a b c u :
+  u = (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ->
+  SN P ->
+  @SN n 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)) /\
@@ -297,6 +371,9 @@ 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.
+
+  move => * /=.
+  apply N_IndSuc';eauto. by asimpl.
 Qed.
 
 Lemma ne_nf_embed n (a : PTm n) :
@@ -336,7 +413,15 @@ Proof.
   move => a b ha iha m ξ[]//= u u0 [? ]. subst.
   case : u0 => //=. move => p p0 [*]. subst.
   spec_refl. by eauto with sn.
-Qed.
+
+  move => P b c ha iha hb ihb hc ihc m ξ []//= 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.
+  case : a0 => //= a0 [*]. subst.
+  spec_refl. eexists; repeat split; eauto with sn.
+  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) /\
@@ -375,7 +460,9 @@ Proof.
       * move => p p0 [*]. subst.
         hauto lq:on db:sn.
   - move => a b ha iha m ρ []//=; first by hauto l:on db:sn.
-    hauto q:on inv:PTm 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 => t0 t1 [*]. subst.
     spec_refl.
@@ -384,6 +471,29 @@ 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 ρ []//=.
+    + 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 ρ []//=.
+    + hauto lq:on db:sn.
+    + move => P0 a0 b0 c0 [?]. subst.
+      case : a0 => //=.
+      * hauto lq:on db:sn.
+      * move => a0 [*]. subst.
+        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 ρ[]//=.
+    + hauto lq:on db:sn.
+    + move => P1 a2 b2 c2 [*]. subst.
+      spec_refl.
+      case : iha.
+      * move => [a3][?]h. subst.
+        left. eexists; split; last by eauto with sn.
+        asimpl. eauto with sn.
+      * hauto lq:on db:sn.
 Qed.
 
 Lemma SN_AppInv : forall n (a b : PTm n), SN (PApp a b) -> SN a /\ SN b.
@@ -978,7 +1088,7 @@ Module Type NoForbid.
   (* 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 PAbs_imp : forall n a b, @ishf n a -> ~~ isabs a -> ~ P (PApp a 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 P_AppInv : forall n (a b : PTm n), P (PApp a b) -> P a /\ P b.
@@ -1023,8 +1133,8 @@ Module SN_NoForbid <: NoForbid.
   Lemma P_RRed : forall n (a b : PTm n), RRed.R a b -> P a -> P b.
   Proof. hauto q:on use:red_sn_preservation, RPar.FromRRed. Qed.
 
-  Lemma PAbs_imp : forall n a b, @ishf n a -> ~~ isabs a -> ~ P (PApp a b).
-    sfirstorder use:fp_red.PAbs_imp. Qed.
+  Lemma PApp_imp : forall n a b, @ishf n 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).
     sfirstorder use:fp_red.PProj_imp. Qed.
 
@@ -1067,7 +1177,7 @@ Module NoForbid_FactSN := NoForbid_Fact SN_NoForbid.
 
 Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
   Import M MFacts.
-  #[local]Hint Resolve P_EPar P_RRed PAbs_imp PProj_imp : forbid.
+  #[local]Hint Resolve P_EPar P_RRed PApp_imp PProj_imp : forbid.
 
   Lemma η_split n (a0 a1 : PTm n) :
     EPar.R a0 a1 ->
@@ -1104,7 +1214,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         have ? : ~~ isabs (ren_PTm shift b) by scongruence use:isabs_ren.
         have ? : ishf (ren_PTm shift b) by scongruence use:ishf_ren.
         exfalso.
-        sfirstorder use:PAbs_imp.
+        sfirstorder use:PApp_imp.
     - move => n a0 a1 h ih /[dup] hP.
       move /P_PairInv => [/P_ProjInv + _].
       move : ih => /[apply].
@@ -1151,7 +1261,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         (* Impossible *)
         * move =>*. exfalso.
           have : P (PApp a2 b0) by sfirstorder use:RReds.AppCong, @rtc_refl, P_RReds.
-          sfirstorder use:PAbs_imp.
+          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 : ih => /[apply]. move => [a2 [iha0 iha1]].
@@ -1231,7 +1341,7 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           have : rtc RRed.R (PApp a0 b0) (PApp (PPair (PProj PL a1) (PProj PR a1)) b0)
             by qauto l:on ctrs:rtc use:RReds.AppCong.
           move : P_RReds hP. repeat move/[apply].
-          sfirstorder use:PAbs_imp.
+          sfirstorder use:PApp_imp.
         * exists (subst_PTm (scons b0 VarPTm) a1).
           split.
           apply : rtc_r; last by apply RRed.AppAbs.

From 2af49373e3fa84b8691a791e1538b60b8c1f61ea Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 21 Feb 2025 15:11:12 -0500
Subject: [PATCH 124/210] Repair epar sn preservation

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

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 7068cb1..9ee4595 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -517,6 +517,14 @@ Proof.
   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.
+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 => //=.
+  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) /\
@@ -527,6 +535,7 @@ Proof.
   - sauto lq:on.
   - sauto lq:on.
   - sauto lq:on.
+  - sauto lq:on.
   - move => a b ha iha hb ihb b0.
     inversion 1; subst.
     + have /iha : (EPar.R (PProj PL a0) (PProj PL b0)) by sauto lq:on.
@@ -545,6 +554,9 @@ Proof.
   - sauto lq:on.
   - sauto lq:on.
   - sauto lq:on.
+  - sauto lq:on.
+  - sauto lq:on.
+  - sauto lq:on.
   - move => a b ha iha c h0.
     inversion h0; subst.
     inversion H1; subst.
@@ -576,6 +588,15 @@ Proof.
       sauto lq:on.
     + sauto lq:on.
   - sauto.
+  - sauto q:on.
+  - move => P a b c hP ihP ha iha hb ihb hc ihc u.
+    elim /EPar.inv => //=_.
+    move => P0 P1 a0 a1 b0 b1 c0 c1 hP0 ha0 hb0 hc0 [*]. subst.
+    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.
+  - sauto q:on.
 Qed.
 
 Module RRed.

From fc0e096c046046a5b182b8d4c5db99a4c2489e96 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 21 Feb 2025 17:27:50 -0500
Subject: [PATCH 125/210] Add two cases for red sn preservation

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

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 9ee4595..41013d1 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -70,17 +70,7 @@ Module EPar.
   | SucCong a0 a1 :
     R a0 a1 ->
     (* ------------ *)
-    R (PSuc a0) (PSuc a1)
-  (* | IndZero P b0 b1  c : *)
-  (*   R b0 b1 -> *)
-  (*   R (PInd P PZero b0 c) b1 *)
-  (* | IndSuc P0 P1 a0 a1 b0 b1 c0 c1 : *)
-  (*   R P0 P1 -> *)
-  (*   R a0 a1 -> *)
-  (*   R b0 b1 -> *)
-  (*   R c0 c1 -> *)
-  (*   (* ----------------------------- *) *)
-  (*   R (PInd P0 (PSuc a0) b0 c0) (subst_PTm (scons (PInd P1 a1 b1 c1) (scons a1 VarPTm)) c1) *).
+    R (PSuc a0) (PSuc a1).
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -601,13 +591,18 @@ Qed.
 
 Module RRed.
   Inductive R {n} : PTm n -> PTm n ->  Prop :=
-  (****************** Eta ***********************)
+  (****************** 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 ********************)
   | AbsCong a0 a1 :
     R a0 a1 ->
@@ -632,7 +627,22 @@ Module RRed.
     R (PBind p A0 B) (PBind p A1 B)
   | BindCong1 p A B0 B1 :
     R B0 B1 ->
-    R (PBind p A B0) (PBind p A B1).
+    R (PBind p A B0) (PBind p A B1)
+  | IndCong0 P0 P1 a b c :
+    R P0 P1 ->
+    R (PInd P0 a b c) (PInd P1 a b c)
+  | IndCong1 P a0 a1 b c :
+    R a0 a1 ->
+    R (PInd P a0 b c) (PInd P a1 b c)
+  | IndCong2 P a b0 b1 c :
+    R b0 b1 ->
+    R (PInd P a b0 c) (PInd P a b1 c)
+  | IndCong3 P a b c0 c1 :
+    R c0 c1 ->
+    R (PInd P a b c0) (PInd P a b c1)
+  | SucCong a0 a1 :
+    R a0 a1 ->
+    R (PSuc a0) (PSuc a1).
 
   Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
@@ -642,15 +652,21 @@ Module RRed.
   Proof.
     move => ->. by apply AppAbs. Qed.
 
+  Lemma IndSuc' n 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.
+  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).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
 
+    all : try qauto ctrs:R.
     move => n a b m ξ /=.
     apply AppAbs'. by asimpl.
-    all : qauto ctrs:R.
+    move => */=; apply IndSuc'; eauto. by asimpl.
   Qed.
 
   Ltac2 rec solve_anti_ren () :=
@@ -672,6 +688,14 @@ Module RRed.
       eexists. split. apply AppAbs. by asimpl.
     - move => n p a b m ξ []//=.
       move => p0 []//=. hauto q:on ctrs:R.
+    - move => n p b c m ξ []//= P a0 b0 c0 [*]. subst.
+      destruct a0 => //=.
+      hauto lq:on ctrs:R.
+    - move => n P a b c m ξ []//=.
+      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) :
@@ -688,9 +712,11 @@ Module RRed.
     R a b -> R (subst_PTm ρ a) (subst_PTm ρ b).
   Proof.
     move => h. move : m ρ. elim : n a b / h => n.
+
+    all : try hauto lq:on ctrs:R.
     move => a b m ρ /=.
     eapply AppAbs'; eauto; cycle 1. by asimpl.
-    all : hauto lq:on ctrs:R.
+    move => */=; apply : IndSuc'; eauto. by asimpl.
   Qed.
 
 End RRed.
@@ -708,6 +734,18 @@ Module RPar.
     R b0 b1 ->
     R (PProj p (PPair a0 b0)) (if p is PL then a1 else b1)
 
+  | IndZero P b0 b1  c :
+    R b0 b1 ->
+    R (PInd P PZero b0 c) b1
+
+  | IndSuc P0 P1 a0 a1 b0 b1 c0 c1 :
+    R P0 P1 ->
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    (* ----------------------------- *)
+    R (PInd P0 (PSuc a0) b0 c0) (subst_PTm (scons (PInd P1 a1 b1 c1) (scons a1 VarPTm)) c1)
+
   (*************** Congruence ********************)
   | AbsCong a0 a1 :
     R a0 a1 ->
@@ -732,7 +770,22 @@ Module RPar.
     R B0 B1 ->
     R (PBind p A0 B0) (PBind p A1 B1)
   | BotCong :
-    R PBot PBot.
+    R PBot PBot
+  | NatCong :
+    R PNat PNat
+  | IndCong P0 P1 a0 a1 b0 b1 c0 c1 :
+    R P0 P1 ->
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    (* ----------------------- *)
+    R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1)
+  | ZeroCong :
+    R PZero PZero
+  | SucCong a0 a1 :
+    R a0 a1 ->
+    (* ------------ *)
+    R (PSuc a0) (PSuc a1).
 
   Lemma refl n (a : PTm n) : R a a.
   Proof.
@@ -755,15 +808,28 @@ Module RPar.
     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 :
+    u = (subst_PTm (scons (PInd P1 a1 b1 c1) (scons a1 VarPTm)) c1) ->
+    R P0 P1 ->
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    (* ----------------------------- *)
+    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) :
     R a b -> R (ren_PTm ξ a) (ren_PTm ξ b).
   Proof.
     move => h. move : m ξ.
     elim : n a b /h.
 
+    all : try qauto ctrs:R use:ProjPair'.
+
     move => n a0 a1 b0 b1 ha iha hb ihb m ξ /=.
     eapply AppAbs'; eauto. by asimpl.
-    all : qauto ctrs:R use:ProjPair'.
+
+    move => * /=. apply : IndSuc'; eauto. by asimpl.
   Qed.
 
   Lemma morphing_ren n m p (ρ0 ρ1 : fin n -> PTm m) (ξ : fin m -> fin p) :
@@ -787,10 +853,13 @@ Module RPar.
     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.
+    all : try hauto lq:on ctrs:R use:morphing_up, ProjPair'.
     move => a0 a1 b0 b1 ha iha hb ihb m ρ0 ρ1 hρ /=.
-    eapply AppAbs'; eauto; cycle 1. sfirstorder use:morphing_up.
+    eapply AppAbs'; eauto; cycle 1. sfirstorder use:morphing_up. by asimpl.
+    move => */=; eapply IndSuc'; eauto; cycle 1.
+    sfirstorder use:morphing_up.
+    sfirstorder use:morphing_up.
     by asimpl.
-    all : hauto lq:on ctrs:R use:morphing_up, ProjPair'.
   Qed.
 
   Lemma substing n m (a : PTm n) b (ρ : fin n -> PTm m) :
@@ -800,7 +869,6 @@ Module RPar.
     hauto l:on use:morphing, refl.
   Qed.
 
-
   Lemma cong n (a0 a1 : PTm (S n)) b0 b1  :
     R a0 a1 ->
     R b0 b1 ->
@@ -828,6 +896,13 @@ Module RPar.
     | 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)
+    | PInd P PZero b c => tstar b
+    | PInd P (PSuc a) b c =>
+        (subst_PTm (scons (PInd (tstar P) (tstar a) (tstar b) (tstar c)) (scons (tstar a) VarPTm)) (tstar c))
+    | PInd P a b c => PInd (tstar P) (tstar a) (tstar b) (tstar c)
     end.
 
   Lemma triangle n (a b : PTm n) :
@@ -859,6 +934,18 @@ Module RPar.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
     - hauto lq:on ctrs:R inv:R.
+    - 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.
+        eauto.
+      + sauto q:on ctrs:R.
+    - sauto lq:on.
   Qed.
 
   Lemma diamond n (a b c : PTm n) :
@@ -876,12 +963,16 @@ Proof.
   - 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 ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN.
   - hauto q:on ctrs:SN inv:SN, TRedSN'.
   - hauto lq:on ctrs:SN inv:RPar.R.
   - hauto lq:on ctrs:SN inv:RPar.R.
+  - hauto l:on inv:RPar.R.
+  - hauto l:on inv:RPar.R.
+  - hauto lq:on ctrs:SN inv:RPar.R.
   - move => a b ha iha hb ihb.
     elim /RPar.inv : ihb => //=_.
     + move => a0 a1 b0 b1 ha0 hb0 [*]. subst.
@@ -901,7 +992,10 @@ Proof.
   - hauto lq:on inv:RPar.R ctrs:RPar.R, TRedSN', TRedSN.
   - hauto lq:on inv:RPar.R ctrs:RPar.R, TRedSN', TRedSN.
   - sauto.
-Qed.
+  - sauto.
+  - admit.
+  - sauto q:on.
+Admitted.
 
 Module RReds.
 

From 9f3b04d0412bfc6883de2d3f09f5cc1ebe261e6a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 21 Feb 2025 22:23:38 -0500
Subject: [PATCH 126/210] Finish sn red preservation

---
 theories/fp_red.v | 14 ++++++++++++--
 1 file changed, 12 insertions(+), 2 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 41013d1..07f5413 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -993,9 +993,19 @@ Proof.
   - hauto lq:on inv:RPar.R ctrs:RPar.R, TRedSN', TRedSN.
   - sauto.
   - sauto.
-  - admit.
+  - move => P a b c hP ihP ha iha hb ihb hc ihc u.
+    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.
+    + 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.
   - sauto q:on.
-Admitted.
+Qed.
 
 Module RReds.
 

From 29d05befe9779426798214425785adf6a076db21 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Fri, 21 Feb 2025 23:43:43 -0500
Subject: [PATCH 127/210] Seemingly redundant nonelim cases

---
 theories/fp_red.v | 95 ++++++++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 93 insertions(+), 2 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 07f5413..ed8354b 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1038,6 +1038,19 @@ Module RReds.
     rtc RRed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
+  Lemma SucCong n (a0 a1 : PTm n) :
+    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 :
+    rtc RRed.R P0 P1 ->
+    rtc RRed.R a0 a1 ->
+    rtc RRed.R b0 b1 ->
+    rtc RRed.R c0 c1 ->
+    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 :
     rtc RRed.R A0 A1 ->
     rtc RRed.R B0 B1 ->
@@ -1053,7 +1066,7 @@ Module RReds.
   Lemma FromRPar n (a b : PTm n) (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.
+    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.
     apply : rtc_r; last by apply RRed.AppAbs.
     by eauto using AppCong, AbsCong.
@@ -1061,6 +1074,12 @@ Module RReds.
     move => n p a0 a1 b0 b1 ha iha hb ihb.
     apply : rtc_r; last by apply RRed.ProjPair.
     by eauto using PairCong, ProjCong.
+
+    hauto lq:on ctrs:RRed.R, rtc.
+
+    move => *.
+    apply : rtc_r; last by apply RRed.IndSuc.
+    by eauto using SucCong, IndCong.
   Qed.
 
   Lemma RParIff n (a b : PTm n) :
@@ -1119,6 +1138,21 @@ Module NeEPar.
     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 :
+    R_nonelim P0 P1 ->
+    R_elim a0 a1 ->
+    R_nonelim b0 b1 ->
+    R_nonelim c0 c1 ->
+    (* ----------------------- *)
+    R_nonelim (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1)
+  | ZeroCong :
+    R_nonelim PZero PZero
+  | SucCong a0 a1 :
+    R_nonelim a0 a1 ->
+    (* ------------ *)
+    R_nonelim (PSuc a0) (PSuc a1)
   with R_elim {n} : PTm n -> PTm n -> Prop :=
   | NAbsCong a0 a1 :
     R_nonelim a0 a1 ->
@@ -1143,7 +1177,22 @@ Module NeEPar.
     R_nonelim B0 B1 ->
     R_elim (PBind p A0 B0) (PBind p A1 B1)
   | NBotCong :
-    R_elim PBot PBot.
+    R_elim PBot PBot
+  | NNatCong :
+    R_elim PNat PNat
+  | NIndCong P0 P1 a0 a1 b0 b1 c0 c1 :
+    R_nonelim P0 P1 ->
+    R_elim a0 a1 ->
+    R_nonelim b0 b1 ->
+    R_nonelim c0 c1 ->
+    (* ----------------------- *)
+    R_elim (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1)
+  | NZeroCong :
+    R_elim PZero PZero
+  | NSucCong a0 a1 :
+    R_nonelim a0 a1 ->
+    (* ------------ *)
+    R_elim (PSuc a0) (PSuc a1).
 
   Scheme epar_elim_ind := Induction for R_elim Sort Prop
       with epar_nonelim_ind := Induction for R_nonelim Sort Prop.
@@ -1162,6 +1211,7 @@ Module NeEPar.
     - hauto lb:on.
     - hauto lq:on inv:R_elim.
     - hauto b:on.
+    - hauto lqb:on inv:R_elim.
     - move => a0 a1 /negP ha' ha ih ha1.
       have {ih} := ih ha1.
       move => ha0.
@@ -1179,6 +1229,7 @@ Module NeEPar.
     - hauto lb: on drew: off.
     - hauto lq:on rew:off inv:R_elim.
     - sfirstorder b:on.
+    - hauto lqb:on inv:R_elim.
   Qed.
 
   Lemma R_nonelim_nothf n (a b : PTm n) :
@@ -1215,10 +1266,15 @@ Module Type NoForbid.
   (* 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,
+  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_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.
@@ -1263,6 +1319,12 @@ Module SN_NoForbid <: NoForbid.
   Lemma PProj_imp : forall  n p a, @ishf n 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,
+      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.
   Proof. sfirstorder use:SN_AppInv. Qed.
 
@@ -1296,6 +1358,9 @@ Module SN_NoForbid <: NoForbid.
   Lemma P_AppBind : forall n p (A : PTm n) 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.
+  Proof. sfirstorder use:SN_IndInv. Qed.
+
 End SN_NoForbid.
 
 Module NoForbid_FactSN := NoForbid_Fact SN_NoForbid.
@@ -1413,6 +1478,32 @@ 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 => []. move : ihP => /[apply].
+      move => [P01][ihP0]ihP1.
+      move => []. move : iha => /[apply].
+      move => [a01][iha0]iha1.
+      move => []. move : ihb => /[apply].
+      move => [b01][ihb0]ihb1.
+      move : ihc => /[apply].
+      move => [c01][ihc0]ihc1.
+      exists
+      case /orP : (orNb (ishf a01)) => [h|].
+      + eexists. split. by eauto using RReds.IndCong.
+        hauto q:on ctrs:NeEPar.R_nonelim use:NeEPar.R_nonelim_nothf.
+      + move => h.
+        case /orP : (orNb (issuc a01 || iszero a01)).
+        * move /norP.
+          have : P (PInd P01 a01 b01 c01) by eauto using P_RReds, RReds.IndCong.
+          hauto lq:on use:PInd_imp.
+        * case /orP.
+          admit.
+          move {h}.
+          case : a01 iha0 iha1 => //=.
+
+best b:on use:PInd_imp.
+
   Qed.
 
 

From d9d0e9cdd4b653115ad2ad6087833a02af65d925 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 22 Feb 2025 00:10:18 -0500
Subject: [PATCH 128/210] One remaining case for eta postponement

---
 theories/fp_red.v | 31 ++++++++++++++++++++++---------
 1 file changed, 22 insertions(+), 9 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index ed8354b..6e9413e 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1273,6 +1273,7 @@ Module Type NoForbid.
 
   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.
 
@@ -1342,6 +1343,9 @@ Module SN_NoForbid <: NoForbid.
     move : p A B E. elim : n u /hu=>//=;sauto lq:on rew:off.
   Qed.
 
+  Lemma P_SucInv : forall n (a : PTm n), P (PSuc a) -> P a.
+  Proof. sauto lq:on. Qed.
+
   Lemma P_AbsInv : forall n (a : PTm (S n)), P (PAbs a) -> P a.
   Proof.
     move => n a. move E : (PAbs a) => u h.
@@ -1488,7 +1492,6 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
       move => [b01][ihb0]ihb1.
       move : ihc => /[apply].
       move => [c01][ihc0]ihc1.
-      exists
       case /orP : (orNb (ishf a01)) => [h|].
       + eexists. split. by eauto using RReds.IndCong.
         hauto q:on ctrs:NeEPar.R_nonelim use:NeEPar.R_nonelim_nothf.
@@ -1497,16 +1500,15 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
         * move /norP.
           have : P (PInd P01 a01 b01 c01) by eauto using P_RReds, RReds.IndCong.
           hauto lq:on use:PInd_imp.
-        * case /orP.
-          admit.
-          move {h}.
-          case : a01 iha0 iha1 => //=.
-
-best b:on use:PInd_imp.
-
+        * move => ha01.
+          eexists. split. eauto using RReds.IndCong.
+          apply NeEPar.IndCong; eauto.
+          move : iha1 ha01. clear.
+          inversion 1; subst => //=; hauto lq:on ctrs:NeEPar.R_elim.
+    - hauto l:on.
+    - hauto lq:on ctrs:NeEPar.R_nonelim use:RReds.SucCong, P_SucInv.
   Qed.
 
-
   Lemma eta_postponement n a b c :
     @P n a ->
     EPar.R a b ->
@@ -1612,6 +1614,17 @@ best b:on use:PInd_imp.
     - 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.
+      admit.
+    - hauto lq:on inv:RRed.R ctrs:rtc, EPar.R.
+    - move => n 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]].
+      exists (PSuc a2).
+      split. by apply RReds.SucCong.
+      by apply EPar.SucCong.
   Qed.
 
   Lemma η_postponement_star n a b c :

From f44c8ef425e88436df3987697e425f10e783d923 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 22 Feb 2025 01:20:35 -0500
Subject: [PATCH 129/210] Only the indsucc case remaining for postponement

---
 theories/fp_red.v | 39 ++++++++++++++++++++++++++++++++++++++-
 1 file changed, 38 insertions(+), 1 deletion(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 6e9413e..ee5492f 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1616,7 +1616,44 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
     - 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.
-      admit.
+      move => /[dup] hInd.
+      move /P_IndInv.
+      move => [pP0][pa0][pb0]pc0.
+      elim /RRed.inv => //= _.
+      + move => P2 b2 c2 [*]. subst.
+        move /η_split : pa0 ha; repeat move/[apply].
+        move => [a1][h0]h1 {iha}.
+        inversion h1; subst.
+        * exfalso.
+          have :P (PInd P0 (PAbs (PApp (ren_PTm shift a2) (VarPTm var_zero))) b0 c0) by eauto using RReds.IndCong, rtc_refl, P_RReds.
+          clear. hauto lq:on use:PInd_imp.
+        * have :P (PInd P0 (PPair (PProj PL a2) (PProj PR a2)) b0 c0) by eauto using RReds.IndCong, rtc_refl, P_RReds.
+          clear. hauto lq:on use:PInd_imp.
+        * eexists. split; eauto.
+          apply : rtc_r.
+          apply : RReds.IndCong; eauto; eauto using rtc_refl.
+          apply RRed.IndZero.
+      + admit.
+      + move => P2 P3 a2 b2 c hP0 [*]. subst.
+        move : ihP hP0 pP0. repeat move/[apply].
+        move => [P2][h0]h1.
+        exists (PInd P2 a0 b0 c0).
+        sfirstorder use:RReds.IndCong, @rtc_refl, EPar.IndCong.
+      + move => P2 a2 a3 b2 c + [*]. subst.
+        move : iha pa0; repeat move/[apply].
+        move => [a2][*].
+        exists (PInd P0 a2 b0 c0).
+        sfirstorder use:RReds.IndCong, @rtc_refl, EPar.IndCong.
+      + move => P2 a2 b2 b3 c + [*]. subst.
+        move : ihb pb0; repeat move/[apply].
+        move => [b2][*].
+        exists (PInd P0 a0 b2 c0).
+        sfirstorder use:RReds.IndCong, @rtc_refl, EPar.IndCong.
+      + move => P2 a2 b2 b3 c + [*]. subst.
+        move : ihc pc0; repeat move/[apply].
+        move => [c2][*].
+        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.
       elim /RRed.inv => //= _ a2 a3 h [*]. subst.

From 2ab47ac883bac2ab760428448bc5eaeec70d308b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 22 Feb 2025 01:29:24 -0500
Subject: [PATCH 130/210] Finish eta postponement

---
 theories/fp_red.v | 14 +++++++++++++-
 1 file changed, 13 insertions(+), 1 deletion(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index ee5492f..2cb864c 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1633,7 +1633,19 @@ Module UniqueNF (M : NoForbid) (MFacts : NoForbid_FactSig M).
           apply : rtc_r.
           apply : RReds.IndCong; eauto; eauto using rtc_refl.
           apply RRed.IndZero.
-      + admit.
+      + move => P2 a2 b2 c [*]. subst.
+        move /η_split /(_ pa0) : ha.
+        move => [a1][h0]h1.
+        inversion h1; subst.
+        * have :P (PInd P0 (PAbs (PApp (ren_PTm shift a3) (VarPTm var_zero))) b0 c0) by eauto using RReds.IndCong, rtc_refl, P_RReds.
+          clear. hauto q:on use:PInd_imp.
+        * have :P (PInd P0 (PPair (PProj PL a3) (PProj PR a3)) b0 c0) by eauto using RReds.IndCong, rtc_refl, P_RReds.
+          clear. hauto q:on use:PInd_imp.
+        * eexists. split.
+          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.
       + move => P2 P3 a2 b2 c hP0 [*]. subst.
         move : ihP hP0 pP0. repeat move/[apply].
         move => [P2][h0]h1.

From 6b8120848bcbd6afa65334183539a643bcacdf8b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 22 Feb 2025 01:33:47 -0500
Subject: [PATCH 131/210] Minor

---
 theories/fp_red.v | 40 ++++++++++++++++++++++++++++++++++++----
 1 file changed, 36 insertions(+), 4 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 2cb864c..62bb50e 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -1743,7 +1743,22 @@ Module ERed.
     R (PBind p A0 B) (PBind p A1 B)
   | BindCong1 p A B0 B1 :
     R B0 B1 ->
-    R (PBind p A B0) (PBind p A B1).
+    R (PBind p A B0) (PBind p A B1)
+  | IndCong0 P0 P1 a b c :
+    R P0 P1 ->
+    R (PInd P0 a b c) (PInd P1 a b c)
+  | IndCong1 P a0 a1 b c :
+    R a0 a1 ->
+    R (PInd P a0 b c) (PInd P a1 b c)
+  | IndCong2 P a b0 b1 c :
+    R b0 b1 ->
+    R (PInd P a b0 c) (PInd P a b1 c)
+  | IndCong3 P a b c0 c1 :
+    R c0 c1 ->
+    R (PInd P a b c0) (PInd P a b c1)
+  | SucCong a0 a1 :
+    R a0 a1 ->
+    R (PSuc a0) (PSuc a1).
 
   Derive Dependent Inversion inv with (forall n (a b : PTm n), R a b) Sort Prop.
 
@@ -1825,7 +1840,11 @@ Module ERed.
       apply. hauto l:on.
     - hauto lq:on rew:off.
     - hauto lq:on rew:off.
-  Qed.
+    - hauto lq:on rew:off.
+    - hauto lq:on rew:off.
+    - hauto lq:on rew:off.
+    - admit.
+  Admitted.
 
   Lemma antirenaming n m (a : PTm n) (b : PTm m) (ξ : fin n -> fin m) :
     (forall i j, ξ i = ξ j -> i = j) ->
@@ -1925,6 +1944,19 @@ Module EReds.
   Proof. solve_s. Qed.
 
 
+  Lemma SucCong n (a0 a1 : PTm n) :
+    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 :
+    rtc ERed.R P0 P1 ->
+    rtc ERed.R a0 a1 ->
+    rtc ERed.R b0 b1 ->
+    rtc ERed.R c0 c1 ->
+    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).
   Proof. induction 1; hauto l:on use:ERed.renaming ctrs:rtc. Qed.
@@ -1933,7 +1965,7 @@ Module EReds.
     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.
+    move => h. elim : n a b /h; eauto using AbsCong, AppCong, PairCong, ProjCong, rtc_refl, BindCong, IndCong, SucCong.
     - move => n 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.
@@ -1976,7 +2008,7 @@ Module EReds.
     move E : (PProj p a) => u hu.
     move : p a E.
     elim : u C /hu;
-      hauto q:on ctrs:rtc,ERed.R inv:ERed.R.
+      scrush ctrs:rtc,ERed.R inv:ERed.R.
   Qed.
 
   Lemma bind_inv n p (A : PTm n) B C :

From f8da81ad7460cef9d74b96d0708e5aa94601aeaa Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 01:13:44 -0500
Subject: [PATCH 132/210] Work on local confluence

---
 theories/fp_red.v | 144 +++++++++++++++++++++++++++++++++++++++++++---
 1 file changed, 136 insertions(+), 8 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 62bb50e..5f01fe2 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2061,6 +2061,12 @@ Module RERed.
   | 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)
+
   (****************** Eta ***********************)
   | AppEta a :
     R (PAbs (PApp (ren_PTm shift a) (VarPTm var_zero))) a
@@ -2091,7 +2097,22 @@ Module RERed.
     R (PBind p A0 B) (PBind p A1 B)
   | BindCong1 p A B0 B1 :
     R B0 B1 ->
-    R (PBind p A B0) (PBind p A B1).
+    R (PBind p A B0) (PBind p A B1)
+  | IndCong0 P0 P1 a b c :
+    R P0 P1 ->
+    R (PInd P0 a b c) (PInd P1 a b c)
+  | IndCong1 P a0 a1 b c :
+    R a0 a1 ->
+    R (PInd P a0 b c) (PInd P a1 b c)
+  | IndCong2 P a b0 b1 c :
+    R b0 b1 ->
+    R (PInd P a b0 c) (PInd P a b1 c)
+  | IndCong3 P a b c0 c1 :
+    R c0 c1 ->
+    R (PInd P a b c0) (PInd P a b c1)
+  | SucCong a0 a1 :
+    R a0 a1 ->
+    R (PSuc a0) (PSuc a1).
 
   Lemma ToBetaEta n (a b : PTm n) :
     R a b ->
@@ -2195,6 +2216,19 @@ Module REReds.
     rtc RERed.R (PProj p a0) (PProj p a1).
   Proof. solve_s. Qed.
 
+  Lemma SucCong n (a0 a1 : PTm n) :
+    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 :
+    rtc RERed.R P0 P1 ->
+    rtc RERed.R a0 a1 ->
+    rtc RERed.R b0 b1 ->
+    rtc RERed.R c0 c1 ->
+    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 :
     rtc RERed.R A0 A1 ->
     rtc RERed.R B0 B1 ->
@@ -2261,8 +2295,22 @@ Module REReds.
   Proof.
     move E : (PProj p a) => u hu.
     move : p a E.
-    elim : u C /hu;
-      hauto q:on ctrs:rtc,RERed.R use:RERed.hne_preservation inv:RERed.R.
+    elim : u C /hu => //=;
+      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) :
+    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 /\
+                     rtc RERed.R a a0 /\
+                     rtc RERed.R b b0 /\
+                     rtc RERed.R c c0.
+  Proof.
+    move E : (PInd P a b c) => u hu.
+    move : P a b c E.
+    elim  : u C / hu => //=;
+      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) :
@@ -2281,12 +2329,17 @@ Module REReds.
          apply rtc_refl.
   Qed.
 
+  Lemma cong_up2 n m (ρ0 ρ1 : fin n -> PTm m) :
+    (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)).
+  Proof. hauto l:on use:cong_up. Qed.
+
   Lemma cong n m (a : PTm n) (ρ0 ρ1 : fin n -> PTm m) :
     (forall i, rtc RERed.R (ρ0 i) (ρ1 i)) ->
     rtc RERed.R (subst_PTm ρ0 a) (subst_PTm ρ1 a).
   Proof.
-    move : m ρ0 ρ1. elim : n / a;
-      eauto using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl.
+    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) :
@@ -2319,6 +2372,13 @@ Module LoRed.
   | 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 ********************)
   | AbsCong a0 a1 :
     R a0 a1 ->
@@ -2348,7 +2408,29 @@ Module LoRed.
   | BindCong1 p A B0 B1 :
     nf A ->
     R B0 B1 ->
-    R (PBind p A B0) (PBind p A B1).
+    R (PBind p A B0) (PBind p A B1)
+  | IndCong0 P0 P1 a b c :
+    ne a ->
+    R P0 P1 ->
+    R (PInd P0 a b c) (PInd P1 a b c)
+  | IndCong1 P a0 a1 b c :
+    ~~ ishf a0 ->
+    R a0 a1 ->
+    R (PInd P a0 b c) (PInd P a1 b c)
+  | IndCong2 P a b0 b1 c :
+    nf P ->
+    ne a ->
+    R b0 b1 ->
+    R (PInd P a b0 c) (PInd P a b1 c)
+  | IndCong3 P a b c0 c1 :
+    nf P ->
+    ne a ->
+    nf b ->
+    R c0 c1 ->
+    R (PInd P a b c0) (PInd P a b c1)
+  | SucCong a0 a1 :
+    R a0 a1 ->
+    R (PSuc a0) (PSuc a1).
 
   Lemma hf_preservation n (a b : PTm n) :
     LoRed.R a b ->
@@ -2368,6 +2450,11 @@ Module LoRed.
     R (PApp (PAbs a) b)  u.
   Proof. move => ->. apply AppAbs. Qed.
 
+  Lemma IndSuc' n 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.
+  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).
   Proof.
@@ -2377,6 +2464,7 @@ Module LoRed.
     move => n a b m ξ /=.
     apply AppAbs'. by asimpl.
     all : try qauto ctrs:R use:ne_nf_ren, ishf_ren.
+    move => * /=; apply IndSuc'. by asimpl.
   Qed.
 
 End LoRed.
@@ -2438,6 +2526,22 @@ Module LoReds.
     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 :
+    rtc LoRed.R a0 a1 ->
+    rtc LoRed.R P0 P1 ->
+    rtc LoRed.R b0 b1 ->
+    rtc LoRed.R c0 c1 ->
+    ne a1 ->
+    nf P1 ->
+    nf b1 ->
+    rtc LoRed.R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
+  Proof. solve_s. Qed.
+
+  Lemma SucCong n (a0 a1 : PTm n) :
+    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,
@@ -2450,12 +2554,16 @@ Module LoReds.
     - 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.
     - sfirstorder use:ne_nf.
     - hauto lq:on ctrs:rtc.
     - hauto lq:on use:LoReds.BindCong solve+:triv.
     - hauto lq:on ctrs:rtc.
+    - hauto lq:on ctrs:rtc.
+    - hauto lq:on ctrs:rtc.
+    - hauto l:on use:SucCong.
     - qauto ctrs:LoRed.R.
     - move => n a0 a1 b hb ihb h.
       have : ~~ ishf a0 by inversion h.
@@ -2465,6 +2573,11 @@ Module LoReds.
     - move => n 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.
+      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.
@@ -2485,6 +2598,10 @@ Fixpoint size_PTm {n} (a : PTm n) :=
   | PPair a b => 3 + Nat.add (size_PTm a) (size_PTm b)
   | 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
+  | PNat => 3
+  | PSuc a => 3 + size_PTm a
+  | PZero => 3
   | PBot => 1
   end.
 
@@ -2575,8 +2692,19 @@ Proof.
       hauto lq:on ctrs:rtc use:EReds.BindCong.
   - move => p A B0 B1 hB ihB u.
     elim /ERed.inv => //=_;
-      hauto lq:on ctrs:rtc use:EReds.BindCong.
-Qed.
+                     hauto lq:on ctrs:rtc use:EReds.BindCong.
+  - move => P0 P1 a b c hP ihP u.
+    elim /ERed.inv => //=_.
+    + move => P2 P3 a0 b0 c0 hP' [*]. subst.
+      move : ihP hP' => /[apply]. move => [P4][hP0]hP1.
+      eauto using EReds.IndCong, rtc_refl.
+    + move => P2 a0 a1 b0 c0  + [*]. subst.
+      move {ihP}.
+      move => h. exists (PInd P1 a1 b c).
+      eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
+    + move => P2 a0 b0 b1 c0 hb [*] {ihP}. subst.
+
+Admitted.
 
 Lemma ered_confluence n (a b c : PTm n) :
   rtc ERed.R a b ->

From ffb57a91f4ac7671e85a13f7c30ac29ca1b902a7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 13:58:35 -0500
Subject: [PATCH 133/210] Update the syntactic reduction lemmas

---
 theories/fp_red.v | 120 +++++++++++++++++++++++++++++++++++++++++-----
 1 file changed, 107 insertions(+), 13 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5f01fe2..5992def 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2022,6 +2022,30 @@ Module EReds.
     hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
+  Lemma suc_inv n (a : PTm n) C :
+    rtc ERed.R (PSuc a) C ->
+    exists b, rtc ERed.R a b /\ C = PSuc b.
+  Proof.
+    move E : (PSuc a) => u hu.
+    move : a E.
+    elim : u C / hu=>//=.
+    - hauto l:on.
+    - 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 :
+    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 /\
+                      C = PInd P0 a0 b0 c0.
+  Proof.
+    move E : (PInd P a b c) => u hu.
+    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.
+  Qed.
+
 End EReds.
 
 #[export]Hint Constructors ERed.R RRed.R EPar.R : red.
@@ -2050,6 +2074,18 @@ Module EJoin.
     hauto lq:on rew:off use:EReds.bind_inv.
   Qed.
 
+  Lemma suc_inj n  (A0 A1 : PTm n)  :
+    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 :
+    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.
+
 End EJoin.
 
 Module RERed.
@@ -2361,6 +2397,17 @@ Module REReds.
     induction h; hauto lq:on rew:off ctrs:rtc use:RERed.ToBetaEta, RReds.nf_refl, @rtc_once, ERed.nf_preservation.
   Qed.
 
+  Lemma suc_inv n (a : PTm n) C :
+    rtc RERed.R (PSuc a) C ->
+    exists b, rtc RERed.R a b /\ C = PSuc b.
+  Proof.
+    move E : (PSuc a) => u hu.
+    move : a E.
+    elim : u C / hu=>//=.
+    - hauto l:on.
+    - hauto lq:on rew:off ctrs:rtc, RERed.R inv:RERed.R, rtc.
+  Qed.
+
 End REReds.
 
 Module LoRed.
@@ -2699,12 +2746,37 @@ Proof.
       move : ihP hP' => /[apply]. move => [P4][hP0]hP1.
       eauto using EReds.IndCong, rtc_refl.
     + move => P2 a0 a1 b0 c0  + [*]. subst.
-      move {ihP}.
-      move => h. exists (PInd P1 a1 b c).
       eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
     + move => P2 a0 b0 b1 c0 hb [*] {ihP}. subst.
-
-Admitted.
+      eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
+    + move => P2 a0 b0 c0 c1 h [*] {ihP}. subst.
+      eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
+  - move => P a0 a1 b c ha iha u.
+    elim /ERed.inv => //=_;
+                     try solve [move => P0 P1 a2 b0 c0 hP[*]; subst;
+      eauto 20 using rtc_refl, EReds.IndCong, rtc_l].
+    move => P0 P1 a2 b0 c0 hP[*]. subst.
+    move : iha hP => /[apply].
+    move => [? [*]].
+    eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
+  - move => P a b0 b1 c hb ihb u.
+    elim /ERed.inv => //=_;
+      try solve [
+          move => P0 P1 a0 b2 c0 hP [*]; subst;
+          eauto 20 using rtc_refl, EReds.IndCong, rtc_l].
+    move => P0 a0 b2 b3 c0 h [*]. subst.
+    move : ihb h => /[apply]. move => [b2 [*]].
+    eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
+  - move => P a b c0 c1 hc ihc u.
+    elim /ERed.inv => //=_;
+      try solve [
+          move => P0 P1 a0 b0 c hP [*]; subst;
+          eauto 20 using rtc_refl, EReds.IndCong, rtc_l].
+    move => P0 a0 b0 c2 c3 h [*]. subst.
+    move : ihc h => /[apply]. move => [c2][*].
+    eauto 20 using rtc_refl, EReds.IndCong, rtc_l.
+  - qauto l:on inv:ERed.R ctrs:rtc use:EReds.SucCong.
+Qed.
 
 Lemma ered_confluence n (a b c : PTm n) :
   rtc ERed.R a b ->
@@ -2820,16 +2892,16 @@ Module BJoin.
     exists v. sfirstorder use:@relations.rtc_transitive.
   Qed.
 
-  Lemma AbsCong n (a b : PTm (S n)) :
-    R a b ->
-    R (PAbs a) (PAbs b).
-  Proof. hauto lq:on use:RReds.AbsCong unfold:R. Qed.
+  (* Lemma AbsCong n (a b : PTm (S n)) : *)
+  (*   R a b -> *)
+  (*   R (PAbs a) (PAbs b). *)
+  (* Proof. hauto lq:on use:RReds.AbsCong unfold:R. Qed. *)
 
-  Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
-    R a0 a1 ->
-    R b0 b1 ->
-    R (PApp a0 b0) (PApp a1 b1).
-  Proof. hauto lq:on use:RReds.AppCong unfold:R. Qed.
+  (* Lemma AppCong n (a0 a1 b0 b1 : PTm n) : *)
+  (*   R a0 a1 -> *)
+  (*   R b0 b1 -> *)
+  (*   R (PApp a0 b0) (PApp a1 b1). *)
+  (* Proof. hauto lq:on use:RReds.AppCong unfold:R. Qed. *)
 End BJoin.
 
 Module DJoin.
@@ -2957,6 +3029,13 @@ Module DJoin.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
+  Lemma suc_inj n  (A0 A1 : PTm n)  :
+    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) :
     R (PApp a0 b0) (PApp a1 b1) ->
     ishne a0 ->
@@ -3251,6 +3330,13 @@ Module ESub.
     sauto lq:on rew:off inv:Sub1.R.
   Qed.
 
+  Lemma suc_inj n (a b : PTm n) :
+    R (PSuc a) (PSuc b) ->
+    R a b.
+  Proof.
+    sauto lq:on use:EReds.suc_inv inv:Sub1.R.
+  Qed.
+
 End ESub.
 
 Module Sub.
@@ -3400,6 +3486,14 @@ Module Sub.
     sauto lq:on rew:off use:REReds.univ_inv.
   Qed.
 
+  Lemma suc_inj n  (A0 A1 : PTm n)  :
+    R (PSuc A0) (PSuc A1) ->
+    R A0 A1.
+  Proof.
+    sauto q:on use:REReds.suc_inv.
+  Qed.
+
+
   Lemma cong n m (a b : PTm (S n)) c d (ρ : fin n -> PTm m) :
     R a b -> DJoin.R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) b).
   Proof.

From bf6d7db877307914b3cee175025409dd8ed0b357 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 14:07:16 -0500
Subject: [PATCH 134/210] Add definition for snat

---
 theories/fp_red.v | 5 +++++
 1 file changed, 5 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5992def..8683b68 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -212,6 +212,11 @@ with TRedSN {n} : PTm n -> PTm n -> Prop :=
 
 Derive Dependent Inversion tred_inv with (forall n (a b : PTm n), TRedSN a b) Sort Prop.
 
+Inductive SNat {n} : PTm n -> 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 ->

From 92e418666f5d04e0fc86cbdf28e632e444b47df7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 14:33:56 -0500
Subject: [PATCH 135/210] Add the case for SNat

---
 theories/logrel.v | 38 ++++++++++++++++++++++++++++++++------
 1 file changed, 32 insertions(+), 6 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 52a4438..88e107a 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -31,6 +31,9 @@ Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop)
   (forall a PB, PF a PB -> ⟦ subst_PTm (scons a VarPTm) B ⟧ i ;; I ↘ PB) ->
   ⟦ PBind p A B ⟧ i ;; I ↘ BindSpace p PA PF
 
+| InterpExt_Nat :
+  ⟦ PNat ⟧ i ;; I ↘ SNat
+
 | InterpExt_Univ j :
   j < i ->
   ⟦ PUniv j ⟧ i ;; I ↘ (I j)
@@ -68,6 +71,7 @@ Proof.
   - hauto q:on ctrs:InterpExt.
   - hauto lq:on rew:off ctrs:InterpExt.
   - hauto q:on ctrs:InterpExt.
+  - hauto q:on ctrs:InterpExt.
   - hauto lq:on ctrs:InterpExt.
 Qed.
 
@@ -88,14 +92,14 @@ Lemma InterpUnivN_nolt n i :
 Proof.
   simp InterpUnivN.
   extensionality A. extensionality PA.
-  set I0 := (fun _ => _).
-  set I1 := (fun _ => _).
   apply InterpExt_lt_eq.
   hauto q:on.
 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)) ->
@@ -107,11 +111,12 @@ Lemma InterpUniv_ind
         (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) ->
         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.
 Proof.
-  move => n P  hSN hBind hUniv hRed.
+  move => n 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.
@@ -144,6 +149,9 @@ Lemma InterpUniv_Step i n A A0 PA :
   ⟦ A ⟧ i ↘ PA.
 Proof. simp InterpUniv. apply InterpExt_Step. Qed.
 
+Lemma InterpUniv_Nat i n :
+  ⟦ PNat : PTm n ⟧ i ↘ SNat.
+Proof. simp InterpUniv. apply InterpExt_Nat. Qed.
 
 #[export]Hint Resolve InterpUniv_Bind InterpUniv_Step InterpUniv_Ne InterpUniv_Univ : InterpUniv.
 
@@ -176,6 +184,14 @@ Proof.
   induction 1; eauto using N_Exp.
 Qed.
 
+Lemma CR_SNat : forall n,  @CR n SNat.
+Proof.
+  move => n. 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 ->
   CR PA /\ SN A.
@@ -202,6 +218,7 @@ Proof.
       apply : N_Exp; eauto using N_β.
       hauto lq:on.
       qauto l:on use:SN_AppInv, SN_NoForbid.P_AbsInv.
+  - qauto use:CR_SNat.
   - hauto l:on ctrs:InterpExt rew:db:InterpUniv.
   - hauto l:on ctrs:SN unfold:CR.
 Qed.
@@ -227,6 +244,7 @@ Proof.
     apply N_AppL => //=.
     hauto q:on use:adequacy.
     + hauto lq:on ctrs:rtc unfold:SumSpace.
+  - hauto lq:on ctrs:SNat.
   - hauto l:on use:InterpUniv_Step.
 Qed.
 
@@ -238,14 +256,14 @@ Proof.
   induction 2; hauto lq:on ctrs:rtc use:InterpUniv_back_clos.
 Qed.
 
-
 Lemma InterpUniv_case n i (A : PTm n) PA :
   ⟦ A ⟧ i ↘ PA ->
-  exists H, rtc TRedSN A H /\  ⟦ H ⟧ i ↘ PA /\ (SNe H \/ isbind H \/ isuniv H).
+  exists H, rtc TRedSN A H /\  ⟦ H ⟧ i ↘ PA /\ (SNe H \/ isbind H \/ isuniv H \/ isnat H).
 Proof.
   move : i A PA. apply InterpUniv_ind => //=.
   hauto lq:on ctrs:rtc use:InterpUniv_Ne.
   hauto l:on use:InterpUniv_Bind.
+  hauto l:on use:InterpUniv_Nat.
   hauto l:on use:InterpUniv_Univ.
   hauto lq:on ctrs:rtc.
 Qed.
@@ -285,6 +303,14 @@ Proof. simp InterpUniv.
        sauto lq:on.
 Qed.
 
+Lemma InterpUniv_Nat_inv n i S :
+  ⟦ PNat : PTm n ⟧ i ↘ S -> S = SNat.
+Proof.
+  simp InterpUniv.
+       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 ->
   S = (fun A => exists PA, ⟦ A ⟧ j ↘ PA) /\ j < i.
@@ -331,7 +357,7 @@ Proof.
       move => [H [/DJoin.FromRedSNs h [h1 h0]]].
       have {h}{}hAB : Sub.R A H by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have {}h0 : SNe H.
-      suff : ~ isbind H /\ ~ isuniv H by itauto.
+      suff : ~ isbind H /\ ~ isuniv H /\ ~ isnat H by itauto.
       move : hA hAB. clear. hauto lq:on db:noconf.
       move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
       tauto.

From fabc39b92aad308d3fbd65b56bb9e9bad2a2ce78 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 14:58:26 -0500
Subject: [PATCH 136/210] Add no confusion lemmas

---
 theories/fp_red.v | 82 +++++++++++++++++++++++++++++++++++++++++++++++
 1 file changed, 82 insertions(+)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 8683b68..0b53d92 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2189,6 +2189,10 @@ Module RERed.
     R a b -> isuniv a -> isuniv b.
   Proof. hauto q:on inv:R. Qed.
 
+  Lemma nat_preservation n (a b : PTm n) :
+    R a b -> isnat a -> isnat b.
+  Proof. hauto q:on inv:R. Qed.
+
   Lemma sne_preservation n (a b : PTm n) :
     R a b -> SNe a -> SNe b.
   Proof.
@@ -2284,6 +2288,10 @@ Module REReds.
     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) :
+    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) :
     rtc RERed.R a b -> SNe a -> SNe b.
   Proof. induction 1; qauto l:on ctrs:rtc use:RERed.sne_preservation. Qed.
@@ -2969,6 +2977,14 @@ Module DJoin.
     sfirstorder use:@rtc_refl unfold:R.
   Qed.
 
+  Lemma sne_nat_noconf n (a b : PTm n) :
+    R a b -> SNe a -> isnat b -> False.
+  Proof.
+    move => [c [? ?]] *.
+    have : SNe c /\ isnat c by sfirstorder use:REReds.sne_preservation, REReds.nat_preservation.
+    qauto l:on inv:SNe.
+  Qed.
+
   Lemma sne_bind_noconf n (a b : PTm n) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
@@ -3016,6 +3032,14 @@ Module DJoin.
     case : c => //=.
   Qed.
 
+  Lemma hne_nat_noconf n (a b : PTm n) :
+    R a b -> ishne a -> isnat b -> False.
+  Proof.
+    move => [c [h0 h1]] h2 h3.
+    have : ishne c /\ isnat c by sfirstorder use:REReds.hne_preservation, REReds.nat_preservation.
+    clear. case : c => //=; itauto.
+  Qed.
+
   Lemma bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
     DJoin.R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
     p0 = p1 /\ DJoin.R A0 A1 /\ DJoin.R B0 B1.
@@ -3404,6 +3428,22 @@ 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) :
+    R a b -> SNe a -> isnat b -> False.
+  Proof.
+    move => [c [d [h0 [h1 h2]]]] *.
+    have : SNe c /\ isnat d by sfirstorder use:REReds.sne_preservation, REReds.nat_preservation, Sub1.sne_preservation.
+    move : h2. clear. hauto q:on inv:Sub1.R, SNe.
+  Qed.
+
+  Lemma nat_sne_noconf n (a b : PTm n) :
+    R a b -> isnat a -> SNe b -> False.
+  Proof.
+    move => [c [d [h0 [h1 h2]]]] *.
+    have : SNe d /\ isnat c by sfirstorder use:REReds.sne_preservation, REReds.nat_preservation.
+    move : h2. clear. hauto q:on inv:Sub1.R, SNe.
+  Qed.
+
   Lemma sne_bind_noconf n (a b : PTm n) :
     R a b -> SNe a -> isbind b -> False.
   Proof.
@@ -3452,6 +3492,48 @@ Module Sub.
     clear. case : c => //=. inversion 2.
   Qed.
 
+  Lemma univ_nat_noconf n (a b : PTm n)  :
+    R a b -> isuniv b -> isnat a -> False.
+  Proof.
+    move => [c[d] [? []]] h0 h1 h2 h3.
+    have : isuniv d by eauto using REReds.univ_preservation.
+    have : isnat c by sfirstorder use:REReds.nat_preservation.
+    inversion h1; subst => //=.
+    clear. case : d => //=.
+  Qed.
+
+  Lemma nat_univ_noconf n (a b : PTm n)  :
+    R a b -> isnat b -> isuniv a -> False.
+  Proof.
+    move => [c[d] [? []]] h0 h1 h2 h3.
+    have : isuniv c by eauto using REReds.univ_preservation.
+    have : isnat d by sfirstorder use:REReds.nat_preservation.
+    inversion h1; subst => //=.
+    clear. case : d => //=.
+  Qed.
+
+  Lemma bind_nat_noconf n (a b : PTm n)  :
+    R a b -> isbind b -> isnat a -> False.
+  Proof.
+    move => [c[d] [? []]] h0 h1 h2 h3.
+    have : isbind d by eauto using REReds.bind_preservation.
+    have : isnat c by sfirstorder use:REReds.nat_preservation.
+    move : h1. clear.
+    inversion 1; subst => //=.
+    case : d h1 => //=.
+  Qed.
+
+  Lemma nat_bind_noconf n (a b : PTm n)  :
+    R a b -> isnat b -> isbind a -> False.
+  Proof.
+    move => [c[d] [? []]] h0 h1 h2 h3.
+    have : isbind c by eauto using REReds.bind_preservation.
+    have : isnat d by sfirstorder use:REReds.nat_preservation.
+    move : h1. clear.
+    inversion 1; subst => //=.
+    case : d h1 => //=.
+  Qed.
+
   Lemma bind_univ_noconf n (a b : PTm n) :
     R a b -> isbind a -> isuniv b -> False.
   Proof.

From 4e9a5582d21d432315a37038886d3f6cf9c1d725 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 15:18:57 -0500
Subject: [PATCH 137/210] Fix InterpUniv Sub

---
 theories/logrel.v | 35 +++++++++++++++++++++++++++++------
 1 file changed, 29 insertions(+), 6 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 88e107a..eec0137 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -280,7 +280,7 @@ Lemma redsns_preservation : forall n a b, @SN n 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: 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 :
   SNe A ->
@@ -367,7 +367,7 @@ Proof.
       move => [H [/DJoin.FromRedSNs h [h1 h0]]].
       have {h}{}hAB : Sub.R H A by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have {}h0 : SNe H.
-      suff : ~ isbind H /\ ~ isuniv H by itauto.
+      suff : ~ isbind H /\ ~ isuniv H /\ ~ isnat H by itauto.
       move : hAB hA h0. clear. hauto lq:on db:noconf.
       move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
       tauto.
@@ -377,7 +377,7 @@ Proof.
       have {hU} {}h : Sub.R (PBind p A B) H
         by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have{h0} : isbind H.
-      suff : ~ SNe H /\ ~ isuniv H by itauto.
+      suff : ~ SNe H /\ ~ isuniv H /\ ~ isnat H by itauto.
       have : isbind (PBind p A B) by scongruence.
       move : h. clear. hauto l:on db:noconf.
       case : H h1 h => //=.
@@ -396,7 +396,7 @@ Proof.
       have {hU} {}h : Sub.R H (PBind p A B)
         by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have{h0} : isbind H.
-      suff : ~ SNe H /\ ~ isuniv H by itauto.
+      suff : ~ SNe H /\ ~ isuniv H /\ ~ isnat H by itauto.
       have : isbind (PBind p A B) by scongruence.
       move : h. clear. move : (PBind p A B). hauto lq:on db:noconf.
       case : H h1 h => //=.
@@ -408,13 +408,36 @@ Proof.
       move => a PB PB' ha hPB hPB'.
       eapply ihPF; eauto.
       apply Sub.cong => //=; eauto using DJoin.refl.
+  - move =>  i B PB h. split.
+    + move => hAB a ha.
+      have ? : SN B by hauto l:on use:adequacy.
+      move /InterpUniv_case : h.
+      move => [H [/DJoin.FromRedSNs h [h1 h0]]].
+      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.
+      move : h0 hAB. clear. qauto l:on db:noconf.
+      case : H h1 hAB h0 => //=.
+      move /InterpUniv_Nat_inv. scongruence.
+    + move => hAB a ha.
+      have ? : SN B by hauto l:on use:adequacy.
+      move /InterpUniv_case : h.
+      move => [H [/DJoin.FromRedSNs h [h1 h0]]].
+      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.
+      move : h0 hAB. clear. qauto l:on db:noconf.
+      case : H h1 hAB h0 => //=.
+      move /InterpUniv_Nat_inv. scongruence.
   - move => i j jlti ih B PB hPB. split.
     + have ? : SN B by hauto l:on use:adequacy.
       move /InterpUniv_case : hPB => [H [/DJoin.FromRedSNs h [h1 h0]]].
       move => hj.
       have {hj}{}h : Sub.R (PUniv j) H by eauto using Sub.transitive, Sub.FromJoin.
       have {h0} : isuniv H.
-      suff : ~ SNe H /\ ~ isbind H by tauto. move : h. clear. hauto lq:on db:noconf.
+      suff : ~ SNe H /\ ~ isbind H /\ ~ isnat H by tauto. move : h. clear. hauto lq:on db:noconf.
       case : H h1 h => //=.
       move => j' hPB h _.
       have {}h : j <= j' by hauto lq:on use: Sub.univ_inj. subst.
@@ -426,7 +449,7 @@ Proof.
       move => hj.
       have {hj}{}h : Sub.R H (PUniv j) by eauto using Sub.transitive, Sub.FromJoin, DJoin.symmetric.
       have {h0} : isuniv H.
-      suff : ~ SNe H /\ ~ isbind H by tauto. move : h. clear. hauto lq:on db:noconf.
+      suff : ~ SNe H /\ ~ isbind H /\ ~ isnat H by tauto. move : h. clear. hauto lq:on db:noconf.
       case : H h1 h => //=.
       move => j' hPB h _.
       have {}h : j' <= j by hauto lq:on use: Sub.univ_inj.

From 8df64ef9893650dfba49e9e77cc07b4707262ec0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sun, 23 Feb 2025 16:01:45 -0500
Subject: [PATCH 138/210] Write down the statement for ST_Ind

---
 theories/logrel.v | 37 +++++++++++++++++++++++++++++++++++++
 1 file changed, 37 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index eec0137..0ac4f8e 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1282,6 +1282,43 @@ Proof.
   apply ST_Univ'. lia.
 Qed.
 
+Lemma ST_Nat n Γ i :
+  Γ ⊨ PNat : PTm n ∈ PUniv i.
+Proof.
+  move => ?. apply SemWt_Univ. move => ρ hρ.
+  eexists. by apply InterpUniv_Nat.
+Qed.
+
+Lemma ST_Zero n Γ :
+  Γ ⊨ PZero : PTm n ∈ PNat.
+Proof.
+  move => ρ hρ. exists 0, SNat. simpl. split.
+  apply InterpUniv_Nat.
+  apply S_Zero.
+Qed.
+
+Lemma ST_Suc n Γ (a : PTm n) :
+  Γ ⊨ a ∈ PNat ->
+  Γ ⊨ PSuc a ∈ PNat.
+Proof.
+  move => ha ρ.
+  move : ha => /[apply] /=.
+  move => [k][PA][h0 h1].
+  move /InterpUniv_Nat_inv : h0 => ?. subst.
+  exists 0, SNat. split. apply InterpUniv_Nat.
+  eauto using S_Suc.
+Qed.
+
+Lemma ST_Ind n Γ P (a : PTm n) b c i :
+  funcomp (ren_PTm shift) (scons 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) ->
+  Γ ⊨ PInd P a b c ∈ subst_PTm (scons a VarPTm) P.
+Proof.
+
+Admitted.
+
 Lemma SSu_Univ n Γ i j :
   i <= j ->
   Γ ⊨ PUniv i : PTm n ≲ PUniv j.

From 5544e401a22334a2f57ffeaf4015a39025f640b8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 24 Feb 2025 01:06:47 -0500
Subject: [PATCH 139/210] Make some progress on the ST_Ind case

---
 theories/logrel.v | 187 +++++++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 186 insertions(+), 1 deletion(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 0ac4f8e..8208c6b 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -683,6 +683,33 @@ Proof.
   hauto q:on solve+:(by asimpl).
 Qed.
 
+(* Definition smorphing_ok {n m} Δ Γ (ρ : fin n -> PTm m) := *)
+(*   forall i ξ k PA, ρ_ok Δ ξ -> Δ *)
+
+(*  Δ ⊨ ρ i ∈ subst_PTm ρ (Γ i). *)
+
+(* Lemma morphing_SemWt : forall n Γ (a A : PTm n), *)
+(*     Γ ⊨ a ∈ A -> forall m Δ (ρ : fin n -> PTm m), *)
+(*       smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ subst_PTm ρ A. *)
+(* Proof. *)
+(*   move => n Γ a A ha m Δ ρ. *)
+(*   rewrite /smorphing_ok => hρ. *)
+(*   move => ξ hξ. *)
+(*   asimpl. *)
+(*   suff : ρ_ok Γ (funcomp (subst_PTm ξ) ρ) by hauto l:on. *)
+(*   move : hξ hρ. clear. *)
+(*   move => hξ hρ i k PA. *)
+(*   specialize (hρ i). *)
+(*   move => h. *)
+(*   replace (funcomp (subst_PTm ξ) ρ i ) with *)
+(*     (subst_PTm ξ (ρ i)); last by asimpl. *)
+(*   move : hρ  hξ => /[apply]. *)
+(*   move => [k0][PA0][h0]h1. *)
+(*   move : h0. asimpl => ?. *)
+(*   have ? : PA0 = PA by eauto using InterpUniv_Functional'.  subst. *)
+(*   done. *)
+(* Qed. *)
+
 Lemma weakening_Sem n Γ (a : PTm n) A B i
   (h0 : Γ ⊨ B ∈ PUniv i)
   (h1 : Γ ⊨ a ∈ A) :
@@ -1309,14 +1336,172 @@ Proof.
   eauto using S_Suc.
 Qed.
 
-Lemma ST_Ind n Γ P (a : PTm n) b c i :
+(* Lemma smorphing_ren n m p Ξ Δ Γ *)
+(*   (ρ : fin n -> PTm m) (ξ : fin m -> fin p) : *)
+(*   renaming_ok Ξ Δ ξ -> smorphing_ok Δ Γ ρ -> *)
+(*   smorphing_ok Ξ Γ (funcomp (ren_PTm ξ) ρ). *)
+(* Proof. *)
+(*   move => hξ hρ. *)
+(*   move => i. *)
+(*   rewrite {1}/funcomp. *)
+(*   have -> : *)
+(*     subst_PTm (funcomp (ren_PTm ξ) ρ) (Γ i) = *)
+(*     ren_PTm ξ (subst_PTm ρ (Γ i)) by asimpl. *)
+(*   eapply renaming_SemWt; eauto. *)
+(* Qed. *)
+
+(* Lemma smorphing_ext n m Δ Γ (ρ : fin n -> PTm m) (a : PTm m) (A : PTm n)  : *)
+(*   smorphing_ok Δ Γ ρ -> *)
+(*   Δ ⊨ a ∈ subst_PTm ρ A -> *)
+(*   smorphing_ok *)
+(*   Δ (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ). *)
+(* Proof. *)
+(*   move => h ha i. destruct i as [i|]; by asimpl. *)
+(* Qed. *)
+
+(* Lemma ρ_ok_smorphing_ok n Γ (ρ : fin n -> PTm 0) : *)
+(*   ρ_ok Γ ρ -> *)
+(*   smorphing_ok null Γ ρ. *)
+(* Proof. *)
+(*   rewrite /ρ_ok /smorphing_ok. *)
+(*   move => h i ξ _. *)
+(*   suff ? : ξ = VarPTm. subst. asimpl. *)
+
+Lemma sn_unmorphing' n : (forall (a : PTm n) (s : SN a), forall m (ρ : fin m -> PTm n) 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).
+  rewrite /up_PTm_PTm.
+  move => h. eapply sn_unmorphing' with (ρ := (scons PBot 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).
+  rewrite /up_PTm_PTm.
+  move => h. eapply sn_unmorphing' with (ρ := (scons PBot (scons PBot VarPTm))); eauto.
+  by asimpl.
+Qed.
+
+Lemma SNat_SN n (a : PTm n) : 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 ->
   Γ ⊨ 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) ->
   Γ ⊨ PInd P a b c ∈ subst_PTm (scons a VarPTm) P.
 Proof.
+  move => hA hc ha hb ρ hρ.
+  move /(_ ρ hρ) : ha => [m][PA][ha0]ha1.
+  move /(_ ρ hρ) : hc => [n][PA0][/InterpUniv_Nat_inv ->].
+  simpl.
+  (* Use localiaztion to block asimpl from simplifying pind *)
+  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 ρ)).
+  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.
+  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.
+
+
+  elim : u /hu.
+  + exists m, PA. split.
+    * move : ha0. by asimpl.
+    * apply : InterpUniv_back_clos; eauto.
+      apply N_IndZero; eauto.
+  + move => a snea.
+    have hρ' : ρ_ok (funcomp (ren_PTm shift) (scons PNat Γ)) (scons a ρ)by
+      apply : ρ_ok_cons; eauto using (InterpUniv_Nat 0); hauto ctrs:SNat.
+    move /SemWt_Univ : (hA) hρ' => /[apply].
+    move => [S0 hS0].
+    exists i, S0. split=>//.
+    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
+      apply : ρ_ok_cons; eauto using (InterpUniv_Nat 0); hauto ctrs:SNat.
+    move /SemWt_Univ : (hA) hρ' => /[apply].
+    move => [S0 hS0].
+    exists i, S0. split => //.
+    apply : InterpUniv_back_clos; eauto.
+    apply N_IndSuc; eauto using SNat_SN.
+    move : (PInd (subst_PTm (up_PTm_PTm ρ) P) a b
+              (subst_PTm (up_PTm_PTm (up_PTm_PTm ρ)) c))  h1.
+    move => r hr.
+    have :
+
+
+
+        move /[swap] => ->.
+    + move => ? a0 ? ih c hc ha. subst.
+      move /(_ a0 ltac:(apply rtc_refl) ha) : ih => [m0][PA1][hPA1]hr.
+      have hρ' : ρ_ok (tNat :: Γ) (a0 .: ρ).
+      {
+        apply : ρ_ok_cons; auto.
+        apply InterpUnivN_Nat.
+        hauto lq:on ctrs:rtc.
+      }
+      have : ρ_ok (A :: tNat :: Γ) ((tInd a[ρ] bs a0) .: (a0 .: ρ))
+        by eauto using ρ_ok_cons.
+      move /hb => {hb} [m1][PA2][hPA2]h.
+      exists m1, PA2.
+      split.
+      * move : hPA2. asimpl.
+        move /InterpUnivN_back_preservation_star. apply.
+        qauto l:on use:Pars_morphing_star,good_Pars_morphing_ext ctrs:rtc.
+      * move : h.
+        move /InterpUnivN_back_clos_star. apply; eauto.
+        subst bs.
+        apply : P_IndSuc_star'; eauto.
+        by asimpl.
+    + move => a0 ? <- _ a1 *.
+      have ? : wne a1 by hauto lq:on.
+      suff  /hA : ρ_ok (tNat :: Γ) (a1 .: ρ).
+      move => [S hS].
+      exists l, S. split=>//.
+      suff ? : wn bs.
+      have ? : wn a[ρ] by sfirstorder use:adequacy.
+      have : wne (tInd a[ρ] bs a1) by auto using wne_ind.
+      eapply adequacy; eauto.
+
+      subst bs.
+      rewrite /SemWt in hb.
+      have /hA : ρ_ok (tNat :: Γ) (var_tm 0 .: ρ).
+      {
+        apply : ρ_ok_cons; auto.
+        apply InterpUnivN_Nat.
+        hauto lq:on ctrs:rtc.
+      }
+      move => [S1 hS1].
+      have /hb : ρ_ok (A :: tNat :: Γ) (var_tm 0 .: (var_tm 0 .: ρ)).
+      {
+        apply : ρ_ok_cons; cycle 2; eauto.
+        apply : ρ_ok_cons; cycle 2; eauto.
+        apply InterpUnivN_Nat.
+        hauto lq:on ctrs:rtc.
+        hauto q:on ctrs:rtc use:adequacy.
+      }
+      move =>[m0][PA1][h1]h2.
+      have : wn b[var_tm 0 .: (var_tm 0 .: ρ)] by hauto q:on use:adequacy.
+      clear => h.
+      apply wn_antirenaming with (ξ :=  var_zero .: (var_zero .: id)).
+      by asimpl.
+
+      apply : ρ_ok_cons; auto.
+      apply InterpUnivN_Nat.
+      hauto lq:on use:adequacy db:nfne.
 Admitted.
 
 Lemma SSu_Univ n Γ i j :

From 2a24700f9a3e866498e8b47d230c2670e1fc45c7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Mon, 24 Feb 2025 01:25:10 -0500
Subject: [PATCH 140/210] One case left for nat

---
 theories/logrel.v | 74 +++++++----------------------------------------
 1 file changed, 10 insertions(+), 64 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 8208c6b..b9294bb 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1415,7 +1415,6 @@ Proof.
   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.
 
-
   elim : u /hu.
   + exists m, PA. split.
     * move : ha0. by asimpl.
@@ -1432,7 +1431,7 @@ Proof.
   + move => a ha [j][S][h0]h1.
     have hρ' : ρ_ok (funcomp (ren_PTm shift) (scons PNat Γ)) (scons (PSuc a) ρ)by
       apply : ρ_ok_cons; eauto using (InterpUniv_Nat 0); hauto ctrs:SNat.
-    move /SemWt_Univ : (hA) hρ' => /[apply].
+    move /SemWt_Univ : (hA) (hρ') => /[apply].
     move => [S0 hS0].
     exists i, S0. split => //.
     apply : InterpUniv_back_clos; eauto.
@@ -1440,68 +1439,15 @@ Proof.
     move : (PInd (subst_PTm (up_PTm_PTm ρ) P) a b
               (subst_PTm (up_PTm_PTm (up_PTm_PTm ρ)) c))  h1.
     move => r hr.
-    have :
-
-
-
-        move /[swap] => ->.
-    + move => ? a0 ? ih c hc ha. subst.
-      move /(_ a0 ltac:(apply rtc_refl) ha) : ih => [m0][PA1][hPA1]hr.
-      have hρ' : ρ_ok (tNat :: Γ) (a0 .: ρ).
-      {
-        apply : ρ_ok_cons; auto.
-        apply InterpUnivN_Nat.
-        hauto lq:on ctrs:rtc.
-      }
-      have : ρ_ok (A :: tNat :: Γ) ((tInd a[ρ] bs a0) .: (a0 .: ρ))
-        by eauto using ρ_ok_cons.
-      move /hb => {hb} [m1][PA2][hPA2]h.
-      exists m1, PA2.
-      split.
-      * move : hPA2. asimpl.
-        move /InterpUnivN_back_preservation_star. apply.
-        qauto l:on use:Pars_morphing_star,good_Pars_morphing_ext ctrs:rtc.
-      * move : h.
-        move /InterpUnivN_back_clos_star. apply; eauto.
-        subst bs.
-        apply : P_IndSuc_star'; eauto.
-        by asimpl.
-    + move => a0 ? <- _ a1 *.
-      have ? : wne a1 by hauto lq:on.
-      suff  /hA : ρ_ok (tNat :: Γ) (a1 .: ρ).
-      move => [S hS].
-      exists l, S. split=>//.
-      suff ? : wn bs.
-      have ? : wn a[ρ] by sfirstorder use:adequacy.
-      have : wne (tInd a[ρ] bs a1) by auto using wne_ind.
-      eapply adequacy; eauto.
-
-      subst bs.
-      rewrite /SemWt in hb.
-      have /hA : ρ_ok (tNat :: Γ) (var_tm 0 .: ρ).
-      {
-        apply : ρ_ok_cons; auto.
-        apply InterpUnivN_Nat.
-        hauto lq:on ctrs:rtc.
-      }
-      move => [S1 hS1].
-      have /hb : ρ_ok (A :: tNat :: Γ) (var_tm 0 .: (var_tm 0 .: ρ)).
-      {
-        apply : ρ_ok_cons; cycle 2; eauto.
-        apply : ρ_ok_cons; cycle 2; eauto.
-        apply InterpUnivN_Nat.
-        hauto lq:on ctrs:rtc.
-        hauto q:on ctrs:rtc use:adequacy.
-      }
-      move =>[m0][PA1][h1]h2.
-      have : wn b[var_tm 0 .: (var_tm 0 .: ρ)] by hauto q:on use:adequacy.
-      clear => h.
-      apply wn_antirenaming with (ξ :=  var_zero .: (var_zero .: id)).
-      by asimpl.
-
-      apply : ρ_ok_cons; auto.
-      apply InterpUnivN_Nat.
-      hauto lq:on use:adequacy db:nfne.
+    have hρ'' :  ρ_ok
+                   (funcomp (ren_PTm shift) (scons P (funcomp (ren_PTm shift) (scons PNat Γ)))) (scons r (scons a ρ)) by
+      eauto using ρ_ok_cons, (InterpUniv_Nat 0).
+    move : hb hρ'' => /[apply].
+    move => [k][PA1][h2]h3.
+    move : h2. asimpl => ?.
+    have ? : PA1 = S0 by  eauto using InterpUniv_Functional'.
+    by subst.
+  +
 Admitted.
 
 Lemma SSu_Univ n Γ i j :

From 9cb7f31cdb26430b43bae47f3bed5365ea59f492 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 24 Feb 2025 13:51:13 -0500
Subject: [PATCH 141/210] Finish ST_Ind

---
 theories/logrel.v | 49 ++++++++++++++++-------------------------------
 1 file changed, 16 insertions(+), 33 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index b9294bb..270843b 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1336,37 +1336,6 @@ Proof.
   eauto using S_Suc.
 Qed.
 
-(* Lemma smorphing_ren n m p Ξ Δ Γ *)
-(*   (ρ : fin n -> PTm m) (ξ : fin m -> fin p) : *)
-(*   renaming_ok Ξ Δ ξ -> smorphing_ok Δ Γ ρ -> *)
-(*   smorphing_ok Ξ Γ (funcomp (ren_PTm ξ) ρ). *)
-(* Proof. *)
-(*   move => hξ hρ. *)
-(*   move => i. *)
-(*   rewrite {1}/funcomp. *)
-(*   have -> : *)
-(*     subst_PTm (funcomp (ren_PTm ξ) ρ) (Γ i) = *)
-(*     ren_PTm ξ (subst_PTm ρ (Γ i)) by asimpl. *)
-(*   eapply renaming_SemWt; eauto. *)
-(* Qed. *)
-
-(* Lemma smorphing_ext n m Δ Γ (ρ : fin n -> PTm m) (a : PTm m) (A : PTm n)  : *)
-(*   smorphing_ok Δ Γ ρ -> *)
-(*   Δ ⊨ a ∈ subst_PTm ρ A -> *)
-(*   smorphing_ok *)
-(*   Δ (funcomp (ren_PTm shift) (scons A Γ)) (scons a ρ). *)
-(* Proof. *)
-(*   move => h ha i. destruct i as [i|]; by asimpl. *)
-(* Qed. *)
-
-(* Lemma ρ_ok_smorphing_ok n Γ (ρ : fin n -> PTm 0) : *)
-(*   ρ_ok Γ ρ -> *)
-(*   smorphing_ok null Γ ρ. *)
-(* Proof. *)
-(*   rewrite /ρ_ok /smorphing_ok. *)
-(*   move => h i ξ _. *)
-(*   suff ? : ξ = VarPTm. subst. asimpl. *)
-
 Lemma sn_unmorphing' n : (forall (a : PTm n) (s : SN a), forall m (ρ : fin m -> PTm n) b, a = subst_PTm ρ b -> SN b).
 Proof. hauto l:on use:sn_unmorphing. Qed.
 
@@ -1447,8 +1416,22 @@ Proof.
     move : h2. asimpl => ?.
     have ? : PA1 = S0 by  eauto using InterpUniv_Functional'.
     by subst.
-  +
-Admitted.
+  + move => a a' hr ha' [k][PA1][h0]h1.
+    have : ρ_ok (funcomp (ren_PTm shift) (scons 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 => //.
+    apply : InterpUniv_back_clos; eauto.
+    apply N_IndCong; eauto.
+    suff : PA1 = PA2 by congruence.
+    move : h0 h2. move : InterpUniv_Join'; repeat move/[apply]. apply.
+    apply DJoin.FromRReds.
+    apply RReds.FromRPar.
+    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.
+Qed.
 
 Lemma SSu_Univ n Γ i j :
   i <= j ->

From 1effbd3d85ffd74486b19c5a6caa634ec24b3f7c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 24 Feb 2025 15:20:30 -0500
Subject: [PATCH 142/210] Finish morphing_SemWt

---
 theories/fp_red.v |  17 +++
 theories/logrel.v | 260 ++++++++++++++++++++++++++++------------------
 2 files changed, 174 insertions(+), 103 deletions(-)

diff --git a/theories/fp_red.v b/theories/fp_red.v
index 0b53d92..cf984a8 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2969,6 +2969,14 @@ Module DJoin.
     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 :
+    R P0 P1 ->
+    R a0 a1 ->
+    R b0 b1 ->
+    R c0 c1 ->
+    R (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1).
+  Proof. hauto q:on use:REReds.IndCong. Qed.
+
   Lemma FromRedSNs n (a b : PTm n) :
     rtc TRedSN a b ->
     R a b.
@@ -3129,6 +3137,15 @@ 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) :
+    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'.
+  Qed.
+
   Lemma pair_inj n (a0 a1 b0 b1 : PTm n) :
     SN (PPair a0 b0) ->
     SN (PPair a1 b1) ->
diff --git a/theories/logrel.v b/theories/logrel.v
index 270843b..bea6f98 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -683,32 +683,56 @@ Proof.
   hauto q:on solve+:(by asimpl).
 Qed.
 
-(* Definition smorphing_ok {n m} Δ Γ (ρ : fin n -> PTm m) := *)
-(*   forall i ξ k PA, ρ_ok Δ ξ -> Δ *)
+Definition smorphing_ok {n m} (Δ : fin m -> PTm m) Γ (ρ : fin n -> PTm m) :=
+  forall ξ, ρ_ok Δ ξ -> ρ_ok Γ (funcomp (subst_PTm ξ) ρ).
 
-(*  Δ ⊨ ρ i ∈ subst_PTm ρ (Γ i). *)
+Lemma smorphing_ren n m p Ξ Δ Γ
+  (ρ : fin n -> PTm m) (ξ : fin m -> fin p) :
+  renaming_ok Ξ Δ ξ -> smorphing_ok Δ Γ ρ ->
+  smorphing_ok Ξ Γ (funcomp (ren_PTm ξ) ρ).
+Proof.
+  move => hξ hρ τ.
+  move /ρ_ok_renaming : hξ => /[apply].
+  move => h.
+  rewrite /smorphing_ok in hρ.
+  asimpl.
+  Check (funcomp τ ξ).
+  set u := funcomp _ _.
+  have : u = funcomp (subst_PTm (funcomp τ ξ)) ρ.
+  subst u. extensionality i. by asimpl.
+  move => ->. by apply hρ.
+Qed.
 
-(* Lemma morphing_SemWt : forall n Γ (a A : PTm n), *)
-(*     Γ ⊨ a ∈ A -> forall m Δ (ρ : fin n -> PTm m), *)
-(*       smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ subst_PTm ρ A. *)
-(* Proof. *)
-(*   move => n Γ a A ha m Δ ρ. *)
-(*   rewrite /smorphing_ok => hρ. *)
-(*   move => ξ hξ. *)
-(*   asimpl. *)
-(*   suff : ρ_ok Γ (funcomp (subst_PTm ξ) ρ) by hauto l:on. *)
-(*   move : hξ hρ. clear. *)
-(*   move => hξ hρ i k PA. *)
-(*   specialize (hρ i). *)
-(*   move => h. *)
-(*   replace (funcomp (subst_PTm ξ) ρ i ) with *)
-(*     (subst_PTm ξ (ρ i)); last by asimpl. *)
-(*   move : hρ  hξ => /[apply]. *)
-(*   move => [k0][PA0][h0]h1. *)
-(*   move : h0. asimpl => ?. *)
-(*   have ? : PA0 = PA by eauto using InterpUniv_Functional'.  subst. *)
-(*   done. *)
-(* Qed. *)
+Lemma smorphing_ext n m Δ Γ (ρ : fin n -> PTm m) (a : PTm m) (A : PTm n)  :
+  smorphing_ok Δ Γ ρ ->
+  Δ ⊨ a ∈ subst_PTm ρ A ->
+  smorphing_ok
+  Δ (funcomp (ren_PTm shift) (scons 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.
+    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),
+      smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ subst_PTm ρ A.
+Proof.
+  move => n Γ a A ha m Δ ρ hρ τ hτ.
+  have {}/hρ {}/ha := hτ.
+  asimpl. eauto.
+Qed.
 
 Lemma weakening_Sem n Γ (a : PTm n) A B i
   (h0 : Γ ⊨ B ∈ PUniv i)
@@ -1229,85 +1253,6 @@ Proof.
   hauto lq:on use: DJoin.cong,  DJoin.ProjCong.
 Qed.
 
-Lemma SE_ProjPair1 n Γ (a b : PTm n) 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.
-Proof.
-  move => h0 h1 h2.
-  apply SemWt_SemEq; eauto using ST_Proj1, ST_Pair.
-  apply DJoin.FromRRed0. apply RRed.ProjPair.
-Qed.
-
-Lemma SE_ProjPair2 n Γ (a b : PTm n) 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.
-Proof.
-  move => h0 h1 h2.
-  apply SemWt_SemEq; eauto using ST_Proj2, ST_Pair.
-  apply : ST_Conv_E. apply : ST_Proj2; eauto. apply : ST_Pair; eauto.
-  hauto l:on use:SBind_inst.
-  apply DJoin.cong. apply DJoin.FromRRed0. apply RRed.ProjPair.
-  apply DJoin.FromRRed0. apply RRed.ProjPair.
-Qed.
-
-Lemma SE_PairEta n Γ (a : PTm n) 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.
-Proof.
-  move => h0 h. apply SemWt_SemEq; eauto.
-  apply : ST_Pair; eauto using ST_Proj1, ST_Proj2.
-  rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
-Qed.
-
-Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) 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.
-Proof.
-  move => hPi.
-  move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha.
-  apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App.
-  apply : ST_Conv_E; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
-Qed.
-
-Lemma SSu_Eq n Γ (A B : PTm n) 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) :
-  Γ ⊨ A ≲ B ->
-  Γ ⊨ B ≲ C ->
-  Γ ⊨ A ≲ C.
-Proof.
-  move => ha hb.
-  apply SemLEq_SemWt in ha, hb.
-  have ? : SN B by hauto l:on use:SemWt_SN.
-  move : ha => [ha0 [i [ha1 ha2]]]. move : hb => [hb0 [j [hb1 hb2]]].
-  qauto l:on use:SemWt_SemLEq, Sub.transitive.
-Qed.
-
-Lemma ST_Univ' n Γ i j :
-  i < j ->
-  Γ ⊨ PUniv i : PTm n ∈ 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).
-Proof.
-  apply ST_Univ'. lia.
-Qed.
 
 Lemma ST_Nat n Γ i :
   Γ ⊨ PNat : PTm n ∈ PUniv i.
@@ -1336,6 +1281,7 @@ Proof.
   eauto using S_Suc.
 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).
 Proof. hauto l:on use:sn_unmorphing. Qed.
 
@@ -1433,6 +1379,114 @@ Proof.
     hauto lq:on inv:option use:RPar.refl.
 Qed.
 
+Lemma SE_SucCong n Γ (a b : PTm n) :
+  Γ ⊨ a ≡ b ∈ PNat ->
+  Γ ⊨ PSuc a ≡ PSuc b ∈ PNat.
+Proof.
+  move /SemEq_SemWt => [ha][hb]he.
+  apply SemWt_SemEq; eauto using ST_Suc.
+  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 ->
+  Γ ⊨ 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) ->
+  Γ ⊨ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0.
+Proof.
+  move /SemEq_SemWt=>[hP0][hP1]hPe.
+  move /SemEq_SemWt=>[ha0][ha1]hae.
+  move /SemEq_SemWt=>[hb0][hb1]hbe.
+  move /SemEq_SemWt=>[hc0][hc1]hce.
+  apply SemWt_SemEq; eauto using ST_Ind, DJoin.IndCong.
+  apply ST_Conv_E with (A := subst_PTm (scons a1 VarPTm) P1) (i := i);
+    last by eauto using DJoin.cong', DJoin.symmetric.
+  apply : ST_Ind; eauto. eapply ST_Conv_E with (i := i); eauto.
+
+
+Lemma SE_ProjPair1 n Γ (a b : PTm n) 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.
+Proof.
+  move => h0 h1 h2.
+  apply SemWt_SemEq; eauto using ST_Proj1, ST_Pair.
+  apply DJoin.FromRRed0. apply RRed.ProjPair.
+Qed.
+
+Lemma SE_ProjPair2 n Γ (a b : PTm n) 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.
+Proof.
+  move => h0 h1 h2.
+  apply SemWt_SemEq; eauto using ST_Proj2, ST_Pair.
+  apply : ST_Conv_E. apply : ST_Proj2; eauto. apply : ST_Pair; eauto.
+  hauto l:on use:SBind_inst.
+  apply DJoin.cong. apply DJoin.FromRRed0. apply RRed.ProjPair.
+  apply DJoin.FromRRed0. apply RRed.ProjPair.
+Qed.
+
+Lemma SE_PairEta n Γ (a : PTm n) 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.
+Proof.
+  move => h0 h. apply SemWt_SemEq; eauto.
+  apply : ST_Pair; eauto using ST_Proj1, ST_Proj2.
+  rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
+Qed.
+
+Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) 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.
+Proof.
+  move => hPi.
+  move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha.
+  apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App.
+  apply : ST_Conv_E; eauto using ST_App, DJoin.cong, DJoin.symmetric, SBind_inst.
+Qed.
+
+Lemma SSu_Eq n Γ (A B : PTm n) 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) :
+  Γ ⊨ A ≲ B ->
+  Γ ⊨ B ≲ C ->
+  Γ ⊨ A ≲ C.
+Proof.
+  move => ha hb.
+  apply SemLEq_SemWt in ha, hb.
+  have ? : SN B by hauto l:on use:SemWt_SN.
+  move : ha => [ha0 [i [ha1 ha2]]]. move : hb => [hb0 [j [hb1 hb2]]].
+  qauto l:on use:SemWt_SemLEq, Sub.transitive.
+Qed.
+
+Lemma ST_Univ' n Γ i j :
+  i < j ->
+  Γ ⊨ PUniv i : PTm n ∈ 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).
+Proof.
+  apply ST_Univ'. lia.
+Qed.
+
 Lemma SSu_Univ n Γ i j :
   i <= j ->
   Γ ⊨ PUniv i : PTm n ≲ PUniv j.

From 133bcd55c27ccc104ff3d345f16ddf4dd54d81a7 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 25 Feb 2025 12:48:42 -0500
Subject: [PATCH 143/210] Need to fix gamma eq

---
 theories/Autosubst2/syntax.v |  4 ++--
 theories/logrel.v            | 13 +++++++++++--
 2 files changed, 13 insertions(+), 4 deletions(-)

diff --git a/theories/Autosubst2/syntax.v b/theories/Autosubst2/syntax.v
index aae3e02..5d04c05 100644
--- a/theories/Autosubst2/syntax.v
+++ b/theories/Autosubst2/syntax.v
@@ -1048,9 +1048,9 @@ Arguments PApp {n_PTm}.
 
 Arguments PAbs {n_PTm}.
 
-#[global] Hint Opaque subst_PTm: rewrite.
+#[global]Hint Opaque subst_PTm: rewrite.
 
-#[global] Hint Opaque ren_PTm: rewrite.
+#[global]Hint Opaque ren_PTm: rewrite.
 
 End Extra.
 
diff --git a/theories/logrel.v b/theories/logrel.v
index bea6f98..5699789 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1388,8 +1388,6 @@ 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 ->
   Γ ⊨ a0 ≡ a1 ∈ PNat ->
@@ -1405,7 +1403,18 @@ Proof.
   apply ST_Conv_E with (A := subst_PTm (scons a1 VarPTm) P1) (i := i);
     last by eauto using DJoin.cong', DJoin.symmetric.
   apply : ST_Ind; eauto. eapply ST_Conv_E with (i := i); eauto.
+  change (PUniv i) with (subst_PTm (scons PZero VarPTm) (@PUniv (S n) i)).
+  apply : morphing_SemWt; eauto.
+  apply smorphing_ext. rewrite /smorphing_ok.
+  move => ξ. rewrite /funcomp. by asimpl.
+  by apply ST_Zero.
+  by apply DJoin.substing. move {hc0}.
+  move => ρ hρ.
+  have : ρ_ok (funcomp (ren_PTm shift) (scons P0 (funcomp (ren_PTm shift) (scons PNat Γ)))) ρ.
+  move   : Γ_eq_ρ_ok hρ => /[apply].
+  apply.
 
+  apply Γ_eq_cons.
 
 Lemma SE_ProjPair1 n Γ (a b : PTm n) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->

From e89aaf14a0fcd05a6d57f0ea82cb999b240fd3e4 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 25 Feb 2025 15:05:08 -0500
Subject: [PATCH 144/210] Define cleaned up version of gamma eq

---
 theories/logrel.v | 81 +++++++++++++++++++++++++++++++++++++++++++++--
 1 file changed, 79 insertions(+), 2 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 5699789..0e66743 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1025,6 +1025,59 @@ Proof.
   hauto l:on use:DJoin.transitive.
 Qed.
 
+Definition Γ_sub' {n} (Γ Δ : fin n -> PTm n) := forall ρ, ρ_ok Δ ρ -> ρ_ok Γ ρ.
+
+Definition Γ_eq' {n} (Γ Δ : fin n -> PTm n) := forall ρ, ρ_ok Δ ρ <-> ρ_ok Γ ρ.
+
+Lemma Γ_sub'_refl n (Γ : fin n -> PTm n) : Γ_sub' Γ Γ.
+  rewrite /Γ_sub'. itauto. Qed.
+
+Lemma Γ_sub'_cons n (Γ Δ : fin n -> PTm n) 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 Δ)).
+Proof.
+  move => hsub hsub' hA hB ρ hρ.
+
+  move => k.
+  move => k0 PA.
+  have : ρ_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.
+  - asimpl.
+    move => h.
+    have {}/hsub' hρ' := hρ_inv.
+    move /SemWt_Univ : (hA) (hρ')=> /[apply].
+    move => [S]hS.
+    move /SemWt_Univ : (hB) (hρ_inv)=>/[apply].
+    move => [S1]hS1.
+    move /(_ None) : hρ (hS1). asimpl => /[apply].
+    suff : forall x, S1 x -> PA x by firstorder.
+    apply : InterpUniv_Sub; eauto.
+    by apply Sub.substing.
+Qed.
+
+Lemma Γ_sub'_SemWt n (Γ Δ : fin n -> PTm n) a A  :
+  Γ_sub' Γ Δ ->
+  Γ ⊨ a ∈ A ->
+  Δ ⊨ a ∈ A.
+Proof.
+  move => hs  ha ρ hρ.
+  have {}/hs hρ' := hρ.
+  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 Δ ρ.
@@ -1042,6 +1095,25 @@ Proof.
   hauto l:on use: DJoin.substing.
 Qed.
 
+Lemma Γ_eq_sub n (Γ Δ : fin n -> PTm n) : Γ_eq' Γ Δ <-> Γ_sub' Γ Δ /\ Γ_sub' Δ Γ.
+Proof. rewrite /Γ_eq' /Γ_sub'. hauto l:on. Qed.
+
+Lemma Γ_eq'_cons n (Γ Δ : fin n -> PTm n) 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 Δ)).
+Proof.
+  move => h.
+  have {h} [h0 h1] : Sub.R A B /\ Sub.R B A by hauto lq:on use:Sub.FromJoin, DJoin.symmetric.
+  repeat rewrite ->Γ_eq_sub.
+  hauto l:on use:Γ_sub'_cons.
+Qed.
+
+Lemma Γ_eq'_refl n (Γ : fin n -> PTm n) : Γ_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 Δ ρ.
@@ -1389,12 +1461,14 @@ Proof.
 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 ->
   Γ ⊨ 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) ->
   Γ ⊨ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0.
 Proof.
+  move => hΓ.
   move /SemEq_SemWt=>[hP0][hP1]hPe.
   move /SemEq_SemWt=>[ha0][ha1]hae.
   move /SemEq_SemWt=>[hb0][hb1]hbe.
@@ -1412,9 +1486,12 @@ Proof.
   move => ρ hρ.
   have : ρ_ok (funcomp (ren_PTm shift) (scons P0 (funcomp (ren_PTm shift) (scons PNat Γ)))) ρ.
   move   : Γ_eq_ρ_ok hρ => /[apply].
-  apply.
+  apply. by eauto using Γ_eq_cons, DJoin.symmetric, DJoin.refl, Γ_eq_refl.
+  apply : SemWff_cons; eauto.
+  apply : SemWff_cons; eauto using (@ST_Nat _ _ 0).
+  move : hc1 => /[apply].
+
 
-  apply Γ_eq_cons.
 
 Lemma SE_ProjPair1 n Γ (a b : PTm n) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->

From 890f97365c9232981bf861ec2d4579867a25b9a8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 25 Feb 2025 15:29:25 -0500
Subject: [PATCH 145/210] Finish the indcong rule

---
 theories/logrel.v | 62 +++++++++++++++++++++++++++++++----------------
 1 file changed, 41 insertions(+), 21 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index 0e66743..ad541a5 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -686,6 +686,10 @@ Qed.
 Definition smorphing_ok {n m} (Δ : fin m -> PTm m) Γ (ρ : fin n -> PTm m) :=
   forall ξ, ρ_ok Δ ξ -> ρ_ok Γ (funcomp (subst_PTm ξ) ρ).
 
+Lemma smorphing_ok_refl n (Δ : fin n -> PTm n) : smorphing_ok Δ Δ VarPTm.
+  rewrite /smorphing_ok => ξ. apply.
+Qed.
+
 Lemma smorphing_ren n m p Ξ Δ Γ
   (ρ : fin n -> PTm m) (ξ : fin m -> fin p) :
   renaming_ok Ξ Δ ξ -> smorphing_ok Δ Γ ρ ->
@@ -724,7 +728,6 @@ Proof.
     eauto using InterpUniv_Functional'.
 Qed.
 
-
 Lemma morphing_SemWt : forall n Γ (a A : PTm n),
     Γ ⊨ a ∈ A -> forall m Δ (ρ : fin n -> PTm m),
       smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ subst_PTm ρ A.
@@ -734,6 +737,16 @@ Proof.
   asimpl. eauto.
 Qed.
 
+Lemma morphing_SemWt_Univ : forall n Γ (a : PTm n) i,
+    Γ ⊨ a ∈ PUniv i -> forall m Δ (ρ : fin n -> PTm m),
+      smorphing_ok Δ Γ ρ -> Δ ⊨ subst_PTm ρ a ∈ PUniv i.
+Proof.
+  move => n Γ a i ha.
+  move => m Δ ρ.
+  have -> : PUniv i = subst_PTm ρ (PUniv i) by reflexivity.
+  by apply morphing_SemWt.
+Qed.
+
 Lemma weakening_Sem n Γ (a : PTm n) A B i
   (h0 : Γ ⊨ B ∈ PUniv i)
   (h1 : Γ ⊨ a ∈ A) :
@@ -783,15 +796,20 @@ Proof.
 Qed.
 
 (* Semantic typing rules *)
+Lemma ST_Var' n Γ (i : fin n) j :
+  Γ ⊨ Γ i ∈ PUniv j ->
+  Γ ⊨ VarPTm i ∈ Γ i.
+Proof.
+  move => /SemWt_Univ h.
+  rewrite /SemWt => ρ /[dup] hρ {}/h [S hS].
+  exists j,S.
+  asimpl. firstorder.
+Qed.
+
 Lemma ST_Var n Γ (i : fin n) :
   ⊨ Γ ->
   Γ ⊨ VarPTm i ∈ Γ i.
-Proof.
-  move /(_ i) => [j /SemWt_Univ h].
-  rewrite /SemWt => ρ /[dup] hρ {}/h [S hS].
-  exists j, S.
-  asimpl. firstorder.
-Qed.
+Proof. hauto l:on use:ST_Var' unfold:SemWff. Qed.
 
 Lemma InterpUniv_Bind_nopf n p i (A : PTm n) B PA :
   ⟦ A ⟧ i ↘ PA ->
@@ -1461,14 +1479,12 @@ Proof.
 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 ->
   Γ ⊨ 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) ->
   Γ ⊨ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0.
 Proof.
-  move => hΓ.
   move /SemEq_SemWt=>[hP0][hP1]hPe.
   move /SemEq_SemWt=>[ha0][ha1]hae.
   move /SemEq_SemWt=>[hb0][hb1]hbe.
@@ -1477,21 +1493,25 @@ Proof.
   apply ST_Conv_E with (A := subst_PTm (scons a1 VarPTm) P1) (i := i);
     last by eauto using DJoin.cong', DJoin.symmetric.
   apply : ST_Ind; eauto. eapply ST_Conv_E with (i := i); eauto.
-  change (PUniv i) with (subst_PTm (scons PZero VarPTm) (@PUniv (S n) i)).
-  apply : morphing_SemWt; eauto.
+  apply : morphing_SemWt_Univ; eauto.
   apply smorphing_ext. rewrite /smorphing_ok.
   move => ξ. rewrite /funcomp. by asimpl.
   by apply ST_Zero.
-  by apply DJoin.substing. move {hc0}.
-  move => ρ hρ.
-  have : ρ_ok (funcomp (ren_PTm shift) (scons P0 (funcomp (ren_PTm shift) (scons PNat Γ)))) ρ.
-  move   : Γ_eq_ρ_ok hρ => /[apply].
-  apply. by eauto using Γ_eq_cons, DJoin.symmetric, DJoin.refl, Γ_eq_refl.
-  apply : SemWff_cons; eauto.
-  apply : SemWff_cons; eauto using (@ST_Nat _ _ 0).
-  move : hc1 => /[apply].
-
-
+  by apply DJoin.substing.
+  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.
+  move /morphing_SemWt. apply. apply smorphing_ext.
+  have -> : (funcomp VarPTm shift) = funcomp (ren_PTm shift) (@VarPTm n) 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 DJoin.renaming. by apply DJoin.substing.
+  apply : morphing_SemWt_Univ; eauto.
+  apply smorphing_ext; eauto using smorphing_ok_refl.
+Qed.
 
 Lemma SE_ProjPair1 n Γ (a b : PTm n) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->

From 4dd2cd7cd8e5f0b7f38a3b0357964f45b5886145 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 25 Feb 2025 15:46:22 -0500
Subject: [PATCH 146/210] Finish indzero and indsuc rules

---
 theories/logrel.v | 32 +++++++++++++++++++++++++++++++-
 theories/typing.v | 19 +++++++++++++++++++
 2 files changed, 50 insertions(+), 1 deletion(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index ad541a5..d01eede 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1513,6 +1513,36 @@ Proof.
   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 ->
+  Γ ⊨ 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) ->
+  Γ ⊨ PInd P PZero b c ≡ b ∈ subst_PTm (scons PZero VarPTm) P.
+Proof.
+  move => hP hb hc.
+  apply SemWt_SemEq; eauto using ST_Zero, ST_Ind.
+  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 ->
+  Γ ⊨ 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) ->
+  Γ ⊨ 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.
+  apply SemWt_SemEq; eauto using ST_Suc, ST_Ind.
+  set Δ := (X in X ⊨ _ ∈ _) in hc.
+  have : smorphing_ok Γ Δ (scons (PInd P a b c) (scons a VarPTm)).
+  apply smorphing_ext. apply smorphing_ext. apply smorphing_ok_refl.
+  done. eauto using ST_Ind.
+  move : morphing_SemWt hc; repeat move/[apply].
+  by asimpl.
+  apply DJoin.FromRRed0.
+  apply RRed.IndSuc.
+Qed.
+
 Lemma SE_ProjPair1 n Γ (a b : PTm n) A B i :
   Γ ⊨ PBind PSig A B ∈ (PUniv i) ->
   Γ ⊨ a ∈ A ->
@@ -1692,4 +1722,4 @@ Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong : sem.
diff --git a/theories/typing.v b/theories/typing.v
index 052eab8..c93e624 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -42,6 +42,25 @@ Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
   ⊢ Γ ->
   Γ ⊢ PUniv i : PTm n ∈ PUniv (S i)
 
+| T_Nat n Γ i :
+  ⊢ Γ ->
+  Γ ⊢ PNat : PTm n ∈ PUniv i
+
+| T_Zero n Γ :
+  ⊢ Γ ->
+  Γ ⊢ PZero : PTm n ∈ PNat
+
+| T_Suc n Γ (a : PTm n) :
+  Γ ⊢ a ∈ PNat ->
+  Γ ⊢ PSuc a ∈ PNat
+
+| T_Ind s Γ P (a : PTm s) b c i :
+  funcomp (ren_PTm shift) (scons 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) ->
+  Γ ⊢ PInd P a b c ∈ subst_PTm (scons a VarPTm) P
+
 | T_Conv n Γ (a : PTm n) A B :
   Γ ⊢ a ∈ A ->
   Γ ⊢ A ≲ B ->

From b2aec9c6ceea20715fc43bb268827d08c1935ecc Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 25 Feb 2025 16:06:04 -0500
Subject: [PATCH 147/210] Add syntactic typing rules

---
 theories/logrel.v |  2 +-
 theories/typing.v | 20 ++++++++++++++++++++
 2 files changed, 21 insertions(+), 1 deletion(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index d01eede..315cb40 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1722,4 +1722,4 @@ Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong SE_IndZero SE_IndSuc : sem.
diff --git a/theories/typing.v b/theories/typing.v
index c93e624..38bdd3f 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -115,6 +115,13 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ 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 ≡ 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) ->
+  Γ ⊢ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0
+
 | E_Conv n Γ (a b : PTm n) A B :
   Γ ⊢ a ≡ b ∈ A ->
   Γ ⊢ A ≲ B ->
@@ -139,6 +146,19 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ 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 ->
+  Γ ⊢ 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) ->
+  Γ ⊢ 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 ->
+  Γ ⊢ 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) ->
+  Γ ⊢ 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 :
   ⊢ Γ ->

From 291d821d94ccc0a5febca9f1898dbec5083bef3f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 25 Feb 2025 16:12:44 -0500
Subject: [PATCH 148/210] Add some admits to work on later

---
 theories/preservation.v |  4 +++-
 theories/structural.v   | 34 ++++++++++++++++++++++++++++------
 2 files changed, 31 insertions(+), 7 deletions(-)

diff --git a/theories/preservation.v b/theories/preservation.v
index 6fe081d..089e7de 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -130,6 +130,8 @@ Proof.
       move : hA0 => /[swap]. move : Su_Transitive. repeat move/[apply].
       move {hS}.
       move => ?. apply : E_Conv; eauto. apply : E_ProjPair2; eauto.
+  - admit.
+  - admit.
   - qauto l:on use:Abs_Inv, E_Conv, regularity_sub0, E_Abs.
   - move => a0 a1 b ha iha Γ A /App_Inv [A0][B0][ih0][ih1]hU.
     have {}/iha iha := ih0.
@@ -170,7 +172,7 @@ Proof.
     have {}/ihA ihA := h1.
     apply : E_Conv; eauto.
     apply E_Bind'; eauto using E_Refl.
-Qed.
+Admitted.
 
 Theorem subject_reduction n Γ (a b A : PTm n) :
   Γ ⊢ a ∈ A ->
diff --git a/theories/structural.v b/theories/structural.v
index 2773c3c..58bafe5 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -92,6 +92,15 @@ Proof.
   move => ->. eauto using T_Pair.
 Qed.
 
+Lemma T_Ind' s Γ P (a : PTm s) b c i U :
+  U = subst_PTm (scons a VarPTm) P ->
+  funcomp (ren_PTm shift) (scons 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) ->
+  Γ ⊢ PInd P a b c ∈ U.
+Proof. move =>->. apply T_Ind. Qed.
+
 Lemma T_Proj2' n Γ (a : PTm n) A B U :
   U = subst_PTm (scons (PProj PL a) VarPTm) B ->
   Γ ⊢ a ∈ PBind PSig A B ->
@@ -103,9 +112,7 @@ Lemma E_Proj2' n Γ i (a b : PTm n) A B U :
   Γ ⊢ 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.
+Proof. move => ->. apply E_Proj2. Qed.
 
 Lemma E_Bind' n Γ i p (A0 A1 : PTm n) B0 B1 :
   Γ ⊢ A0 ∈ PUniv i ->
@@ -184,6 +191,7 @@ Proof.
   - move => n Γ a A b B i hA ihA hB ihB hS ihS m Δ ξ 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.
+  - admit.
   - hauto lq:on ctrs:Wt,LEq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on rew:off use:E_Bind', Wff_Cons, renaming_up.
@@ -199,6 +207,7 @@ Proof.
     move : ihb hΔ hξ. repeat move/[apply].
     by asimpl.
   - move => *. apply : E_Proj2'; eauto. by asimpl.
+  - admit.
   - qauto l:on ctrs:Eq, LEq.
   - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ξ hΔ hξ.
     move : ihP (hξ) (hΔ). repeat move/[apply].
@@ -216,6 +225,8 @@ Proof.
   - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ξ hΔ hξ.
     apply : E_ProjPair2'; eauto. by asimpl.
     move : ihb hξ hΔ; repeat move/[apply]. by asimpl.
+  - admit.
+  - admit.
   - move => *.
     apply : E_AppEta'; eauto. by asimpl.
   - qauto l:on use:E_PairEta.
@@ -228,7 +239,7 @@ Proof.
   - qauto l:on ctrs:LEq.
   - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
-Qed.
+Admitted.
 
 Definition morphing_ok {n m} Δ Γ (ρ : fin n -> PTm m) :=
   forall i, Δ ⊢ ρ i ∈ subst_PTm ρ (Γ i).
@@ -342,6 +353,10 @@ Proof.
   - move => *. apply : T_Proj2'; eauto. by asimpl.
   - hauto lq:on ctrs:Wt,LEq.
   - qauto l:on ctrs:Wt.
+  - qauto l:on ctrs:Wt.
+  - qauto l:on ctrs:Wt.
+  - admit.
+  - qauto l:on ctrs:Wt.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on ctrs:Eq.
@@ -359,6 +374,7 @@ Proof.
     by asimpl.
   - hauto q:on ctrs:Eq.
   - move => *. apply : E_Proj2'; eauto. by asimpl.
+  - admit.
   - qauto l:on ctrs:Eq, LEq.
   - move => n Γ a b A B i hP ihP hb ihb ha iha m Δ ρ hΔ hρ.
     move : ihP (hρ) (hΔ). repeat move/[apply].
@@ -376,6 +392,8 @@ Proof.
   - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ρ hΔ hρ.
     apply : E_ProjPair2'; eauto. by asimpl.
     move : ihb hρ hΔ; repeat move/[apply]. by asimpl.
+  - admit.
+  - admit.
   - move => *.
     apply : E_AppEta'; eauto. by asimpl.
   - qauto l:on use:E_PairEta.
@@ -388,7 +406,7 @@ Proof.
   - qauto l:on ctrs:LEq.
   - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
-Qed.
+Admitted.
 
 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.
 Proof. sfirstorder use:morphing. Qed.
@@ -505,6 +523,7 @@ Proof.
     exists j. have : Γ ⊢ PProj PL a ∈ A by qauto use:T_Proj1.
     move : substing_wt h1; repeat move/[apply].
     by asimpl.
+  - admit.
   - sfirstorder.
   - sfirstorder.
   - sfirstorder.
@@ -535,9 +554,12 @@ Proof.
     eauto using bind_inst.
     move /T_Proj1 in iha.
     hauto lq:on ctrs:Wt,Eq,LEq use:Bind_Inv, substing_wt.
+  - admit.
   - hauto lq:on ctrs:Wt.
   - hauto q:on use:substing_wt db:wt.
   - hauto l:on use:bind_inst db:wt.
+  - admit.
+  - admit.
   - move => n Γ b A B i hΓ ihΓ 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)
@@ -603,7 +625,7 @@ Proof.
     + apply Cumulativity with (i := i1); eauto.
       have : Γ ⊢ a1 ∈ A1 by eauto using T_Conv.
       move : substing_wt ih1';repeat move/[apply]. by asimpl.
-Qed.
+Admitted.
 
 Lemma Var_Inv n Γ (i : fin n) A :
   Γ ⊢ VarPTm i ∈ A ->

From cc0e9219c4766fb06289fce1644fefaa05446d86 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 25 Feb 2025 16:50:12 -0500
Subject: [PATCH 149/210] Finish two cases of renaming

---
 theories/structural.v | 58 +++++++++++++++++++++++++++++++++++++------
 1 file changed, 50 insertions(+), 8 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index 58bafe5..09c0b03 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -12,6 +12,11 @@ Proof. apply wt_mutual; eauto. Qed.
 
 #[export]Hint Constructors Wt Wff Eq : wt.
 
+Lemma T_Nat' n Γ :
+  ⊢ Γ ->
+  Γ ⊢ PNat : PTm n ∈ PUniv 0.
+Proof. apply T_Nat. Qed.
+
 Lemma renaming_up n m (ξ : fin n -> fin m) Δ Γ A :
   renaming_ok Δ Γ ξ ->
   renaming_ok (funcomp (ren_PTm shift) (scons (ren_PTm ξ A) Δ)) (funcomp (ren_PTm shift) (scons A Γ)) (upRen_PTm_PTm ξ) .
@@ -169,6 +174,20 @@ 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 :
+  U = subst_PTm (scons PZero VarPTm) P ->
+  funcomp (ren_PTm shift) (scons 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) ->
+  Γ ⊢ PInd P PZero b c ≡ b ∈ U.
+Proof. move => ->. apply E_IndZero. Qed.
+
+Lemma Wff_Cons' n Γ (A : PTm n) i :
+  Γ ⊢ A ∈ PUniv i ->
+  (* -------------------------------- *)
+  ⊢ funcomp (ren_PTm shift) (scons A Γ).
+Proof. hauto lq:on rew:off use:Wff_Cons, wff_mutual. 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 Δ Γ ξ ->
@@ -191,7 +210,27 @@ Proof.
   - move => n Γ a A b B i hA ihA hB ihB hS ihS m Δ ξ 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.
-  - admit.
+  - move => s Γ P a b c i hP ihP ha iha hb ihb hc ihc m Δ ξ 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 _ _.
+    have  hΞ : ⊢ Ξ. subst Ξ.
+    apply : Wff_Cons'; eauto. apply hP.
+    move /(_ _ Ξ (upRen_PTm_PTm (upRen_PTm_PTm ξ)) hΞ) : ihc.
+    set Ξ' := (funcomp _ _ : fin (S (S _)) -> _) .
+    have : renaming_ok Ξ Ξ' (upRen_PTm_PTm (upRen_PTm_PTm ξ)).
+    by do 2 apply renaming_up.
+    move /[swap] /[apply].
+    set T0 := (X in Ξ ⊢ _ ∈ X).
+    set T1 := (Y in _ -> Ξ ⊢ _ ∈ Y).
+    (* Report an autosubst bug *)
+    suff : T0 = T1 by congruence.
+    subst T0 T1. clear.
+    asimpl. substify. asimpl. rewrite /funcomp. asimpl. rewrite /funcomp. done.
   - hauto lq:on ctrs:Wt,LEq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on rew:off use:E_Bind', Wff_Cons, renaming_up.
@@ -225,7 +264,16 @@ Proof.
   - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ξ hΔ hξ.
     apply : E_ProjPair2'; eauto. by asimpl.
     move : ihb hξ hΔ; repeat move/[apply]. by asimpl.
-  - admit.
+  - move => n Γ P i b c hP ihP hb ihb hc ihc m Δ ξ 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.
+    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)).
+    asimpl. rewrite /funcomp. asimpl. apply.
   - admit.
   - move => *.
     apply : E_AppEta'; eauto. by asimpl.
@@ -302,12 +350,6 @@ Proof.
   apply : T_Var';eauto. rewrite /funcomp. by asimpl.
 Qed.
 
-Lemma Wff_Cons' n Γ (A : PTm n) i :
-  Γ ⊢ A ∈ PUniv i ->
-  (* -------------------------------- *)
-  ⊢ funcomp (ren_PTm shift) (scons A Γ).
-Proof. hauto lq:on rew:off use:Wff_Cons, wff_mutual. Qed.
-
 Lemma weakening_wt : forall n Γ (a A B : PTm n) i,
     Γ ⊢ B ∈ PUniv i ->
     Γ ⊢ a ∈ A ->

From 96bc223b0a54b853205ee23e5d2baf4c98cf6883 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 25 Feb 2025 20:18:40 -0500
Subject: [PATCH 150/210] Finish renaming and preservation

---
 theories/common.v     |   6 +-
 theories/structural.v | 156 ++++++++++++++++++++++++++++++++++++------
 theories/typing.v     |   1 +
 3 files changed, 142 insertions(+), 21 deletions(-)

diff --git a/theories/common.v b/theories/common.v
index 9cce0cc..282038d 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.fintype Autosubst2.syntax ssreflect.
+Require Import Autosubst2.fintype Autosubst2.syntax Autosubst2.core ssreflect.
 From Ltac2 Require Ltac2.
 Import Ltac2.Notations.
 Import Ltac2.Control.
@@ -106,3 +106,7 @@ Proof. case : a => //=. Qed.
 Definition ishne_ren n m (a : PTm n)  (ξ : fin n -> fin m) :
   ishne (ren_PTm ξ a) = ishne a.
 Proof. move : m ξ. elim : n / 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.
diff --git a/theories/structural.v b/theories/structural.v
index 09c0b03..b294887 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -50,7 +50,10 @@ Proof.
   move E :(PBind p A B) => T h.
   move : p A B E.
   elim : n Γ T U / h => //=.
-  - hauto lq:on ctrs:Wt,LEq,Eq use:Wt_Univ.
+  - move => n Γ 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.
 
@@ -97,6 +100,16 @@ Proof.
   move => ->. eauto using T_Pair.
 Qed.
 
+Lemma E_IndCong' n Γ P0 P1 (a0 a1 : PTm n) 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 ->
+  Γ ⊢ 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) ->
+  Γ ⊢ 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 :
   U = subst_PTm (scons a VarPTm) P ->
   funcomp (ren_PTm shift) (scons PNat Γ) ⊢ P ∈ PUniv i ->
@@ -188,6 +201,16 @@ Lemma Wff_Cons' n Γ (A : PTm n) i :
   ⊢ funcomp (ren_PTm shift) (scons 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 :
+  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 ->
+  Γ ⊢ 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) ->
+  Γ ⊢ 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 Δ Γ ξ ->
@@ -225,12 +248,7 @@ Proof.
     have : renaming_ok Ξ Ξ' (upRen_PTm_PTm (upRen_PTm_PTm ξ)).
     by do 2 apply renaming_up.
     move /[swap] /[apply].
-    set T0 := (X in Ξ ⊢ _ ∈ X).
-    set T1 := (Y in _ -> Ξ ⊢ _ ∈ Y).
-    (* Report an autosubst bug *)
-    suff : T0 = T1 by congruence.
-    subst T0 T1. clear.
-    asimpl. substify. asimpl. rewrite /funcomp. asimpl. rewrite /funcomp. done.
+    by asimpl.
   - hauto lq:on ctrs:Wt,LEq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on rew:off use:E_Bind', Wff_Cons, renaming_up.
@@ -246,7 +264,27 @@ Proof.
     move : ihb hΔ hξ. repeat move/[apply].
     by asimpl.
   - move => *. apply : E_Proj2'; eauto. by asimpl.
-  - admit.
+  - 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).
+    by asimpl.
+    abstract : hP'.
+    qauto l:on use:renaming_up, Wff_Cons', T_Nat'.
+    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 _ _.
+    have hΞ : ⊢ Ξ.
+    subst Ξ. apply :Wff_Cons'; eauto. apply hP'.
+    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
+         (subst_PTm
+            (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift)) P0)) = ren_PTm shift
+      (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift))
+         (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 : ihP (hξ) (hΔ). repeat move/[apply].
@@ -273,8 +311,23 @@ Proof.
     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)).
-    asimpl. rewrite /funcomp. asimpl. apply.
-  - admit.
+    suff : ren_PTm (upRen_PTm_PTm (upRen_PTm_PTm ξ))
+      (ren_PTm shift
+         (subst_PTm
+            (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift)) P))= ren_PTm shift
+      (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ξ.
+    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 _ _.
+    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)).
+    by asimpl.
   - move => *.
     apply : E_AppEta'; eauto. by asimpl.
   - qauto l:on use:E_PairEta.
@@ -287,7 +340,7 @@ Proof.
   - qauto l:on ctrs:LEq.
   - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
-Admitted.
+Qed.
 
 Definition morphing_ok {n m} Δ Γ (ρ : fin n -> PTm m) :=
   forall i, Δ ⊢ ρ i ∈ subst_PTm ρ (Γ i).
@@ -331,10 +384,6 @@ Lemma renaming_wt' : forall n m Δ Γ a A (ξ : fin n -> fin m) u U,
     renaming_ok Δ Γ ξ -> Δ ⊢ u ∈ U.
 Proof. hauto use:renaming_wt. 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 morphing_up n m Γ Δ (ρ : fin n -> PTm m) (A : PTm n) k :
   morphing_ok Γ Δ ρ ->
   Γ ⊢ subst_PTm ρ A ∈ PUniv k ->
@@ -397,7 +446,29 @@ Proof.
   - qauto l:on ctrs:Wt.
   - qauto l:on ctrs:Wt.
   - qauto l:on ctrs:Wt.
-  - admit.
+  - 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 /morphing_up : hξ.  move /(_ PNat 0).
+    apply. by apply T_Nat.
+    move => [:hP].
+    apply : T_Ind'; eauto. by asimpl.
+    abstract :hP. apply ihP. by eauto using Wff_Cons', T_Nat'.
+    move /morphing_up : hξ.  move /(_ PNat 0).
+    apply. by apply T_Nat.
+    move : ihb hξ hΔ; do!move/[apply]. by asimpl.
+    set Ξ := funcomp _ _.
+    have  hΞ : ⊢ Ξ. subst Ξ.
+    apply : Wff_Cons'; eauto. apply hP.
+    move /(_ _ Ξ (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    set Ξ' := (funcomp _ _ : fin (S (S _)) -> _) .
+    have : morphing_ok Ξ Ξ' (up_PTm_PTm (up_PTm_PTm ξ)).
+    eapply morphing_up; eauto. apply hP.
+    move /[swap]/[apply].
+    set x := (x in _ ⊢ _ ∈ x).
+    set y := (y in _ -> _ ⊢ _ ∈ y).
+    suff : x = y by scongruence.
+    subst x y. asimpl. substify. by asimpl.
   - qauto l:on ctrs:Wt.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on ctrs:Eq.
@@ -416,7 +487,26 @@ Proof.
     by asimpl.
   - hauto q:on ctrs:Eq.
   - move => *. apply : E_Proj2'; eauto. by asimpl.
-  - admit.
+  - 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 /morphing_up : hξ.  move /(_ PNat 0).
+    apply. by apply T_Nat.
+    move => [:hP'].
+    apply E_IndCong' with (i := i).
+    by asimpl.
+    abstract : hP'.
+    qauto l:on use:morphing_up, Wff_Cons', T_Nat'.
+    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 _ _.
+    have hΞ : ⊢ Ξ.
+    subst Ξ. apply :Wff_Cons'; eauto. apply hP'.
+    move /(_ _ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    move /(_ ltac:(qauto l:on use:morphing_up)).
+    asimpl. substify. 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 : ihP (hρ) (hΔ). repeat move/[apply].
@@ -434,8 +524,34 @@ Proof.
   - move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ρ hΔ hρ.
     apply : E_ProjPair2'; eauto. by asimpl.
     move : ihb hρ hΔ; repeat move/[apply]. by asimpl.
-  - admit.
-  - admit.
+  - 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 /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.
+    asimpl. by apply.
+    set Ξ := funcomp _ _.
+    have hΞ : ⊢ Ξ by  qauto l:on use:Wff_Cons', T_Nat', renaming_up.
+    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 /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 _ _.
+    have hΞ : ⊢ Ξ by  qauto l:on use:Wff_Cons', T_Nat', renaming_up.
+    move /(_ _ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
+    move /(_ ltac:(sauto l:on use:morphing_up)).
+    asimpl. substify. by asimpl.
   - move => *.
     apply : E_AppEta'; eauto. by asimpl.
   - qauto l:on use:E_PairEta.
@@ -448,7 +564,7 @@ Proof.
   - qauto l:on ctrs:LEq.
   - move => *; apply : Su_Pi_Proj2'; eauto; by asimpl.
   - move => *. apply : Su_Sig_Proj2'; eauto; by asimpl.
-Admitted.
+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.
 Proof. sfirstorder use:morphing. Qed.
diff --git a/theories/typing.v b/theories/typing.v
index 38bdd3f..2d08215 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -116,6 +116,7 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   Γ ⊢ 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 ->
   Γ ⊢ a0 ≡ a1 ∈ PNat ->
   Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons PZero VarPTm) P0 ->

From bb39f157baaf7976a95d0b8d8f3b2220a49e9427 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 25 Feb 2025 21:04:32 -0500
Subject: [PATCH 151/210] Finish regularity

---
 theories/structural.v | 47 +++++++++++++++++++++++++++++++++++++------
 1 file changed, 41 insertions(+), 6 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index b294887..3fc7ffd 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -542,7 +542,7 @@ Proof.
     have  hξ' : morphing_ok (funcomp (ren_PTm shift) (scons PNat Δ))
                   (funcomp (ren_PTm shift) (scons PNat Γ)) (up_PTm_PTm ξ).
     move /morphing_up : hξ.  move /(_ PNat 0).
-    apply. by apply T_Nat.
+    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.
@@ -681,7 +681,8 @@ Proof.
     exists j. have : Γ ⊢ PProj PL a ∈ A by qauto use:T_Proj1.
     move : substing_wt h1; repeat move/[apply].
     by asimpl.
-  - admit.
+  - move => s Γ P a b c i + ? + *. clear. move => h ha. exists i.
+    hauto lq:on use:substing_wt.
   - sfirstorder.
   - sfirstorder.
   - sfirstorder.
@@ -712,12 +713,45 @@ Proof.
     eauto using bind_inst.
     move /T_Proj1 in iha.
     hauto lq:on ctrs:Wt,Eq,LEq use:Bind_Inv, substing_wt.
-  - admit.
+  - move => n Γ P0 P1 a0 a1 b0 b1 c0 c1 i _ _ hPE [hP0 [hP1 _]] hae [ha0 [ha1 _]] _ [hb0 [hb1 hb2]] _ [hc0 [hc1 hc2]].
+    have wfΓ : ⊢ Γ by hauto use:wff_mutual.
+    repeat split. by eauto using T_Ind.
+    apply : T_Conv. apply : T_Ind; eauto.
+    apply : T_Conv; eauto.
+    eapply morphing; by eauto using Su_Eq, morphing_ext, morphing_id with wt.
+    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.
+    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.
+    sauto lq:on use:substing_wt.
   - hauto lq:on ctrs:Wt.
   - hauto q:on use:substing_wt db:wt.
   - hauto l:on use:bind_inst db:wt.
-  - admit.
-  - admit.
+  - move => n Γ 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 _.
+    have ? : ⊢ Γ by hauto use:wff_mutual.
+    repeat split=>//.
+    apply : T_Ind; eauto with wt.
+    set Ξ : fin (S (S _)) -> _ := (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].
     repeat split => //=; eauto with wt.
     have {}hb : funcomp (ren_PTm shift) (scons A Γ) ⊢ ren_PTm shift b ∈ ren_PTm shift (PBind PPi A B)
@@ -783,7 +817,8 @@ Proof.
     + apply Cumulativity with (i := i1); eauto.
       have : Γ ⊢ a1 ∈ A1 by eauto using T_Conv.
       move : substing_wt ih1';repeat move/[apply]. by asimpl.
-Admitted.
+      Unshelve. all: exact 0.
+Qed.
 
 Lemma Var_Inv n Γ (i : fin n) A :
   Γ ⊢ VarPTm i ∈ A ->

From 687d1be03f8b204f711814b19b32a2b00e0192ba Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Tue, 25 Feb 2025 22:35:59 -0500
Subject: [PATCH 152/210] Finish preservation

---
 theories/algorithmic.v  | 11 +++++++++-
 theories/fp_red.v       |  5 +++++
 theories/logrel.v       | 11 +++++++++-
 theories/preservation.v | 46 ++++++++++++++++++++++++++++++++++++++---
 theories/structural.v   |  2 ++
 theories/typing.v       |  4 ++++
 6 files changed, 74 insertions(+), 5 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 6fe5fad..73a8dc6 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -16,13 +16,22 @@ Module HRed.
   | 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).
+    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 n (a b : PTm n) : HRed.R a b -> RRed.R a b.
   Proof. induction 1; hauto lq:on ctrs:RRed.R. Qed.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index cf984a8..5ad793b 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2977,6 +2977,11 @@ 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) :
+    R a0 a1 ->
+    R (PSuc a0) (PSuc a1).
+  Proof. qauto l:on use:REReds.SucCong. Qed.
+
   Lemma FromRedSNs n (a b : PTm n) :
     rtc TRedSN a b ->
     R a b.
diff --git a/theories/logrel.v b/theories/logrel.v
index 315cb40..a245362 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1578,6 +1578,15 @@ Proof.
   rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
 Qed.
 
+Lemma SE_Nat n Γ (a b : PTm n) :
+  Γ ⊨ a ≡ b ∈ PNat ->
+  Γ ⊨ PSuc a ≡ PSuc b ∈ PNat.
+Proof.
+  move /SemEq_SemWt => [ha][hb]hE.
+  apply SemWt_SemEq; eauto using ST_Suc.
+  eauto using DJoin.SucCong.
+Qed.
+
 Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) A B :
   Γ ⊨ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊨ b0 ≡ b1 ∈ PBind PPi A B ->
@@ -1722,4 +1731,4 @@ Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong SE_IndZero SE_IndSuc : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong SE_IndZero SE_IndSuc SE_SucCong : sem.
diff --git a/theories/preservation.v b/theories/preservation.v
index 089e7de..b78f87e 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -76,6 +76,23 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
+Lemma Ind_Inv n Γ P (a : PTm n) b c U :
+  Γ ⊢ PInd P a b c ∈ U ->
+  exists i, funcomp (ren_PTm shift) (scons 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) /\
+       Γ ⊢ 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 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.
+    eauto using E_Refl, Su_Eq.
+  - 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),
   Γ ⊢ PApp (PAbs a) b ∈ A -> Γ ⊢ PApp (PAbs a) b ≡ subst_PTm (scons b VarPTm) a ∈ A.
 Proof.
@@ -108,6 +125,19 @@ Proof.
   apply : E_ProjPair1; eauto.
 Qed.
 
+Lemma Suc_Inv n Γ (a : PTm n) 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.
+    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 :
   Γ ⊢ a ∈ A ->
   RRed.R a b ->
@@ -130,8 +160,13 @@ Proof.
       move : hA0 => /[swap]. move : Su_Transitive. repeat move/[apply].
       move {hS}.
       move => ?. apply : E_Conv; eauto. apply : E_ProjPair2; eauto.
-  - admit.
-  - admit.
+  - hauto lq:on use:Ind_Inv, E_Conv, E_IndZero.
+  - move => P a b c Γ A.
+    move /Ind_Inv.
+    move => [i][hP][ha][hb][hc]hSu.
+    apply : E_Conv; eauto.
+    apply : E_IndSuc'; eauto.
+    hauto l:on use:Suc_Inv.
   - qauto l:on use:Abs_Inv, E_Conv, regularity_sub0, E_Abs.
   - move => a0 a1 b ha iha Γ A /App_Inv [A0][B0][ih0][ih1]hU.
     have {}/iha iha := ih0.
@@ -172,7 +207,12 @@ Proof.
     have {}/ihA ihA := h1.
     apply : E_Conv; eauto.
     apply E_Bind'; eauto using E_Refl.
-Admitted.
+  - hauto lq:on rew:off use:Ind_Inv, E_Conv, E_IndCong db:wt.
+  - hauto lq:on rew:off use:Ind_Inv, E_Conv, E_IndCong db:wt.
+  - hauto lq:on rew:off use:Ind_Inv, E_Conv, E_IndCong db:wt.
+  - hauto lq:on rew:off use:Ind_Inv, E_Conv, E_IndCong db:wt.
+  - hauto lq:on use:Suc_Inv, E_Conv, E_SucCong.
+Qed.
 
 Theorem subject_reduction n Γ (a b A : PTm n) :
   Γ ⊢ a ∈ A ->
diff --git a/theories/structural.v b/theories/structural.v
index 3fc7ffd..c25986c 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -507,6 +507,7 @@ Proof.
     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 : ihP (hρ) (hΔ). repeat move/[apply].
@@ -734,6 +735,7 @@ Proof.
     by eauto using Su_Sig_Proj2.
     sauto lq:on use:substing_wt.
   - hauto lq:on ctrs:Wt.
+  - 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 _.
diff --git a/theories/typing.v b/theories/typing.v
index 2d08215..818d6b5 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -123,6 +123,10 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
   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) ->
   Γ ⊢ PInd P0 a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P0
 
+| E_SucCong n Γ (a b : PTm n) :
+  Γ ⊢ a ≡ b ∈ PNat ->
+  Γ ⊢ PSuc a ≡ PSuc b ∈ PNat
+
 | E_Conv n Γ (a b : PTm n) A B :
   Γ ⊢ a ≡ b ∈ A ->
   Γ ⊢ A ≲ B ->

From 2a492a67deb2aa3da0f732aa80b2bc3f6e980c6e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 26 Feb 2025 00:46:11 -0500
Subject: [PATCH 153/210] Add algorithmic rules for nat

---
 theories/algorithmic.v | 81 ++++++++++++++++++++++++++++++++++++++++--
 1 file changed, 78 insertions(+), 3 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 73a8dc6..a5a7d0a 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -91,6 +91,17 @@ Lemma T_Bot_Imp n Γ (A : PTm n) :
   induction hu => //=.
 Qed.
 
+Lemma Zero_Inv n Γ U :
+  Γ ⊢ @PZero n ∈ U ->
+  Γ ⊢ PNat ≲ U.
+Proof.
+  move E : PZero => u hu.
+  move : E.
+  elim : n Γ 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 :
   Γ ⊢ PBind p A B ≲ C ->
   exists i A0 B0, Γ ⊢ C ≡ PBind p A0 B0 ∈ PUniv i.
@@ -130,6 +141,21 @@ Proof.
     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.
+    have {}h : Γ ⊢ PZero ∈ PUniv i by hauto l:on use:regularity.
+    exfalso. move : h. clear.
+    move /Zero_Inv /synsub_to_usub.
+    hauto l:on use:Sub.univ_nat_noconf.
+  - move => a h.
+    have {}h : Γ ⊢ PSuc a  ∈ PUniv i by hauto l:on use:regularity.
+    exfalso. move /Suc_Inv : h => [_ /synsub_to_usub].
+    hauto lq:on use:Sub.univ_nat_noconf.
+  - move => P0 a0 b0 c0 h0 h1 /synsub_to_usub [_ [_ h2]].
+    set u := PInd _ _ _ _ in h0.
+    have hne : SNe u by sfirstorder use:ne_nf_embed.
+    exfalso. move : h2 hne. hauto l:on use:Sub.bind_sne_noconf.
 Qed.
 
 Lemma Sub_Univ_InvR n (Γ : fin n -> PTm n) i C :
@@ -163,6 +189,20 @@ Proof.
   - 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.
+    exfalso. move : h. clear.
+    move /Zero_Inv /synsub_to_usub.
+    hauto l:on use:Sub.univ_nat_noconf.
+  - move => a h.
+    have {}h : Γ ⊢ PSuc a  ∈ PUniv j by hauto l:on use:regularity.
+    exfalso. move /Suc_Inv : h => [_ /synsub_to_usub].
+    hauto lq:on use:Sub.univ_nat_noconf.
+  - move => P0 a0 b0 c0 h0 h1 /synsub_to_usub [_ [_ h2]].
+    set u := PInd _ _ _ _ in h0.
+    have hne : SNe u by sfirstorder use:ne_nf_embed.
+    exfalso. move : h2 hne. hauto l:on use:Sub.univ_sne_noconf.
 Qed.
 
 Lemma Sub_Bind_InvL n Γ p (A : PTm n) B C :
@@ -204,6 +244,22 @@ Proof.
     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.
+    have {}h : Γ ⊢ PZero ∈ PUniv i by hauto l:on use:regularity.
+    exfalso. move : h. clear.
+    move /Zero_Inv /synsub_to_usub.
+    hauto l:on use:Sub.univ_nat_noconf.
+  - move => a h.
+    have {}h : Γ ⊢ PSuc a  ∈ PUniv i by hauto l:on use:regularity.
+    exfalso. move /Suc_Inv : h => [_ /synsub_to_usub].
+    hauto lq:on use:Sub.univ_nat_noconf.
+  - move => P0 a0 b0 c0 h0 h1 /synsub_to_usub [_ [_ h2]].
+    set u := PInd _ _ _ _ in h0.
+    have hne : SNe u by sfirstorder use:ne_nf_embed.
+    exfalso. move : h2 hne. subst u.
+    hauto l:on use:Sub.sne_bind_noconf.
 Qed.
 
 Lemma T_Abs_Inv n Γ (a0 a1 : PTm (S n)) U :
@@ -236,6 +292,14 @@ 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 :=
+| CE_ZeroZero :
+  PZero ↔ PZero
+
+| CE_SucSuc a b :
+  a ⇔ b ->
+  (* ------------- *)
+  PSuc a ↔ PSuc b
+
 | CE_AbsAbs a b :
   a ⇔ b ->
   (* --------------------- *)
@@ -283,6 +347,10 @@ Inductive CoqEq {n} : PTm n -> PTm n -> Prop :=
   (* ---------------------------- *)
   PBind p A0 B0 ↔ PBind p A1 B1
 
+| CE_NatCong :
+  (* ------------------ *)
+  PNat ↔ PNat
+
 | CE_NeuNeu a0 a1 :
   a0 ∼ a1 ->
   a0 ↔ a1
@@ -307,6 +375,16 @@ with CoqEq_Neu  {n} : PTm n -> PTm n -> Prop :=
   (* ------------------------- *)
   PApp u0 a0 ∼ PApp u1 a1
 
+| CE_IndCong P0 P1 u0 u1 b0 b1 c0 c1 :
+  ishne u0 ->
+  ishne u1 ->
+  P0 ⇔ P1 ->
+  u0 ∼ u1 ->
+  b0 ⇔ b1 ->
+  c0 ⇔ c1 ->
+  (* ----------------------------------- *)
+  PInd P0 u0 b0 c0 ∼ PInd P1 u1 b1 c1
+
 with CoqEq_R {n} : PTm n -> PTm n -> Prop :=
 | CE_HRed a a' b b'  :
   rtc HRed.R a a' ->
@@ -337,9 +415,6 @@ Lemma coqeq_symmetric_mutual : forall n,
     (forall (a b : PTm n), a ⇔ b -> b ⇔ a).
 Proof. apply coqeq_mutual; qauto l:on ctrs:CoqEq,CoqEq_R, CoqEq_Neu. Qed.
 
-
-(* Lemma Sub_Univ_InvR *)
-
 Lemma coqeq_sound_mutual : forall n,
     (forall (a b : PTm n), a ∼ b -> forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> exists C,
        Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ≡ b ∈ C) /\

From 49bb2cca13047118e35520de8c6f0d4d17d1e0cf Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 26 Feb 2025 15:45:40 -0500
Subject: [PATCH 154/210] Finish the completeness proof

---
 theories/algorithmic.v | 138 +++++++++++++++++++++++++----------------
 1 file changed, 86 insertions(+), 52 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index a5a7d0a..a7dd223 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -287,6 +287,60 @@ Proof.
     apply : ctx_eq_subst_one; eauto.
 Qed.
 
+Lemma T_Univ_Raise n Γ (a : PTm n) 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 :
+  Γ ⊢ PBind p A B ∈ PUniv i ->
+  Γ ⊢ A ∈ PUniv i /\ funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i.
+Proof.
+  move /Bind_Inv.
+  move => [i0][hA][hB]h.
+  move /synsub_to_usub : h => [_ [_ /Sub.univ_inj ? ]].
+  sfirstorder use:T_Univ_Raise.
+Qed.
+
+Lemma Abs_Pi_Inv n Γ (a : PTm (S n)) A B :
+  Γ ⊢ PAbs a ∈ PBind PPi A B ->
+  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B.
+Proof.
+  move => h.
+  have [i hi] : exists i, Γ ⊢ PBind PPi A B ∈ PUniv i by hauto use:regularity.
+  have  [{}i {}hi] : exists i, Γ ⊢ A ∈ PUniv i by hauto use:Bind_Inv.
+  apply : subject_reduction; last apply RRed.AppAbs'.
+  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.
+  by eauto using Wff_Cons'.
+  rewrite -/ren_PTm.
+  by asimpl.
+  rewrite -/ren_PTm.
+  by asimpl.
+Qed.
+
+Lemma Pair_Sig_Inv n Γ (a b : PTm n) A B :
+  Γ ⊢ PPair a b ∈ PBind PSig A B ->
+  Γ ⊢ a ∈ A /\ Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B.
+Proof.
+  move => /[dup] h0 h1.
+  have [i hr] : exists i, Γ ⊢ PBind PSig A B ∈ PUniv i by sfirstorder use:regularity.
+  move /T_Proj1 in h0.
+  move /T_Proj2 in h1.
+  split.
+  hauto lq:on use:subject_reduction ctrs:RRed.R.
+  have hE : Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A by
+    hauto lq:on use:RRed_Eq ctrs:RRed.R.
+  apply : T_Conv.
+  move /subject_reduction : h1. apply.
+  apply RRed.ProjPair.
+  apply : bind_inst; eauto.
+Qed.
+
+
 (* Coquand's algorithm with subtyping *)
 Reserved Notation "a ∼ b" (at level 70).
 Reserved Notation "a ↔ b" (at level 70).
@@ -485,6 +539,37 @@ 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 /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'.
+    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.
+    abstract : sigeq. hauto lq:on use:T_Univ_Raise solve+:lia.
+    have sigleq : Γ ⊢ PBind PSig PNat P0 ≲ PBind PSig PNat P1.
+    apply Su_Sig with (i := 0)=>//. by apply T_Nat'. by eauto using Su_Eq, T_Nat', E_Refl.
+    apply Su_Eq with (i := max i0 i1).  apply sigeq.
+    exists (subst_PTm (scons u0 VarPTm) P0). repeat split => //.
+    suff : Γ ⊢ subst_PTm (scons u0 VarPTm) P0 ≲ subst_PTm (scons u1 VarPTm) P1 by eauto using Su_Transitive.
+    by eauto using Su_Sig_Proj2.
+    apply E_IndCong with (i := max i0 i1); eauto. move :sigeq; clear; hauto q:on use:regularity.
+    apply ihb; eauto. apply : T_Conv; eauto. eapply morphing. apply : Su_Eq. apply E_Symmetric. apply sigeq.
+    done. apply morphing_ext. apply morphing_id. done. by apply T_Zero.
+    apply ihc; eauto.
+    eapply T_Conv; eauto.
+    eapply ctx_eq_subst_one; eauto. apply : Su_Eq; apply sigeq.
+    eapply weakening_su; eauto.
+    eapply morphing; eauto. apply : Su_Eq. apply E_Symmetric. apply sigeq.
+    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.
+  - 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 /Abs_Inv : h0 => [A0][B0][h0]h0'.
     move /Abs_Inv : h1 => [A1][B1][h1]h1'.
@@ -617,6 +702,7 @@ Proof.
     apply : ctx_eq_subst_one; eauto. apply : Su_Eq; apply eqA.
     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.
     have [*] : Γ ⊢ a' ∈ A /\ Γ ⊢ b' ∈ A by eauto using HReds.preservation.
     hauto lq:on use:HReds.ToEq, E_Symmetric, E_Transitive.
@@ -1050,58 +1136,6 @@ Proof.
   - exists i1,j1,b2,b3. sfirstorder b:on solve+:lia.
 Qed.
 
-Lemma T_Univ_Raise n Γ (a : PTm n) 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 :
-  Γ ⊢ PBind p A B ∈ PUniv i ->
-  Γ ⊢ A ∈ PUniv i /\ funcomp (ren_PTm shift) (scons A Γ) ⊢ B ∈ PUniv i.
-Proof.
-  move /Bind_Inv.
-  move => [i0][hA][hB]h.
-  move /synsub_to_usub : h => [_ [_ /Sub.univ_inj ? ]].
-  sfirstorder use:T_Univ_Raise.
-Qed.
-
-Lemma Abs_Pi_Inv n Γ (a : PTm (S n)) A B :
-  Γ ⊢ PAbs a ∈ PBind PPi A B ->
-  funcomp (ren_PTm shift) (scons A Γ) ⊢ a ∈ B.
-Proof.
-  move => h.
-  have [i hi] : exists i, Γ ⊢ PBind PPi A B ∈ PUniv i by hauto use:regularity.
-  have  [{}i {}hi] : exists i, Γ ⊢ A ∈ PUniv i by hauto use:Bind_Inv.
-  apply : subject_reduction; last apply RRed.AppAbs'.
-  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.
-  by eauto using Wff_Cons'.
-  rewrite -/ren_PTm.
-  by asimpl.
-  rewrite -/ren_PTm.
-  by asimpl.
-Qed.
-
-Lemma Pair_Sig_Inv n Γ (a b : PTm n) A B :
-  Γ ⊢ PPair a b ∈ PBind PSig A B ->
-  Γ ⊢ a ∈ A /\ Γ ⊢ b ∈ subst_PTm (scons a VarPTm) B.
-Proof.
-  move => /[dup] h0 h1.
-  have [i hr] : exists i, Γ ⊢ PBind PSig A B ∈ PUniv i by sfirstorder use:regularity.
-  move /T_Proj1 in h0.
-  move /T_Proj2 in h1.
-  split.
-  hauto lq:on use:subject_reduction ctrs:RRed.R.
-  have hE : Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A by
-    hauto lq:on use:RRed_Eq ctrs:RRed.R.
-  apply : T_Conv.
-  move /subject_reduction : h1. apply.
-  apply RRed.ProjPair.
-  apply : bind_inst; eauto.
-Qed.
 
 Lemma coqeq_complete' n k (a b : PTm n) :
   algo_metric k a b ->

From f8943e3a9ca8c24aa6f38b024d46ef01550b51eb Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 26 Feb 2025 16:49:02 -0500
Subject: [PATCH 155/210] Finish some cases of the soundness proof

---
 theories/algorithmic.v | 145 ++++++++++++++++++++++++++++++++++++++---
 1 file changed, 137 insertions(+), 8 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index a7dd223..d7e226c 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -723,8 +723,14 @@ Proof.
   - hauto l:on use:HRed.AppAbs.
   - hauto l:on use:HRed.ProjPair.
   - hauto lq:on ctrs:HRed.R.
+  - hauto lq:on ctrs:HRed.R.
+  - hauto lq:on ctrs:HRed.R.
   - sfirstorder use:ne_hne.
   - hauto lq:on ctrs:HRed.R.
+  - sfirstorder use:ne_hne.
+  - hauto q:on ctrs:HRed.R.
+  - hauto lq:on use:ne_hne.
+  - hauto lq:on use:ne_hne.
 Qed.
 
 Lemma algo_metric_case n k (a b : PTm n) :
@@ -744,6 +750,15 @@ Proof.
     inversion E; subst => /=.
     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + sfirstorder use:ne_hne.
 Qed.
 
 Lemma algo_metric_sym n k (a b : PTm n) :
@@ -779,6 +794,53 @@ Proof.
   clear. move /synsub_to_usub. hauto l:on use:Sub.bind_inj.
 Qed.
 
+Lemma T_AbsZero_Imp n Γ a (A : PTm n) :
+  Γ ⊢ PAbs a ∈ A ->
+  Γ ⊢ PZero ∈ A ->
+  False.
+Proof.
+  move /Abs_Inv => [A0][B0][_]haU.
+  move /Zero_Inv => hbU.
+  move /Sub_Bind_InvR : haU => [i][A2][B2]h2.
+  have : Γ ⊢ PNat ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
+  clear. hauto lq:on use:synsub_to_usub, Sub.bind_nat_noconf.
+Qed.
+
+Lemma T_AbsSuc_Imp n Γ a b (A : PTm n) :
+  Γ ⊢ PAbs a ∈ A ->
+  Γ ⊢ PSuc b ∈ A ->
+  False.
+Proof.
+  move /Abs_Inv => [A0][B0][_]haU.
+  move /Suc_Inv => [_ hbU].
+  move /Sub_Bind_InvR : haU => [i][A2][B2]h2.
+  have {hbU h2} : Γ ⊢ PNat ≲ PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
+  hauto lq:on use:Sub.bind_nat_noconf, synsub_to_usub.
+Qed.
+
+Lemma Nat_Inv n Γ A:
+  Γ ⊢ @PNat n ∈ A ->
+  exists i, Γ ⊢ PUniv i ≲ A.
+Proof.
+  move E : PNat => u hu.
+  move : E.
+  elim : n Γ 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) :
+  Γ ⊢ PAbs a ∈ A ->
+  Γ ⊢ PNat ∈ A ->
+  False.
+Proof.
+  move /Abs_Inv => [A0][B0][_]haU.
+  move /Nat_Inv => [i hA].
+  move /Sub_Univ_InvR : hA => [j][k]hA.
+  have : Γ ⊢ PBind PPi A0 B0 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
+  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 :
   Γ ⊢ PPair a0 a1 ∈ U ->
   Γ ⊢ PBind p A0 B0  ∈ U ->
@@ -791,6 +853,39 @@ 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 :
+  Γ ⊢ PPair a0 a1 ∈ U ->
+  Γ ⊢ PNat ∈ U ->
+  False.
+Proof.
+  move/Pair_Inv => [A1][B1][_][_]hbU.
+  move /Nat_Inv => [i]/Sub_Univ_InvR[j][k]hU.
+  have : Γ ⊢ PBind PSig A1 B1 ≲ PUniv j by eauto using Su_Transitive, Su_Eq.
+  clear. move /synsub_to_usub. hauto l:on use:Sub.bind_univ_noconf.
+Qed.
+
+Lemma T_PairZero_Imp n Γ (a0 a1 : PTm n) U :
+  Γ ⊢ PPair a0 a1 ∈ U ->
+  Γ ⊢ PZero ∈ U ->
+  False.
+Proof.
+  move/Pair_Inv=>[A1][B1][_][_]hbU.
+  move/Zero_Inv. move/Sub_Bind_InvR : hbU=>[i][A0][B0]*.
+  have : Γ ⊢ PNat ≲ PBind PSig A0 B0 by eauto using Su_Transitive, Su_Eq.
+  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 :
+  Γ ⊢ PPair a0 a1 ∈ U ->
+  Γ ⊢ PSuc b ∈ U ->
+  False.
+Proof.
+  move/Pair_Inv=>[A1][B1][_][_]hbU.
+  move/Suc_Inv=>[_ hU]. move/Sub_Bind_InvR : hbU=>[i][A0][B0]*.
+  have : Γ ⊢ PNat ≲ PBind PSig A0 B0 by eauto using Su_Transitive, Su_Eq.
+  clear. move /synsub_to_usub. hauto l:on use:Sub.bind_nat_noconf.
+Qed.
+
 Lemma Univ_Inv n Γ i U :
   Γ ⊢ @PUniv n i ∈ U ->
   Γ ⊢ PUniv i ∈ PUniv (S i)  /\ Γ ⊢ PUniv (S i) ≲ U.
@@ -859,7 +954,7 @@ Proof.
   move : a E.
   elim : u b /hu.
   - hauto l:on.
-  - hauto lq:on ctrs:nsteps inv:LoRed.R.
+  - scrush ctrs:nsteps inv:LoRed.R.
 Qed.
 
 Lemma lored_hne_preservation n (a b : PTm n) :
@@ -1116,7 +1211,6 @@ Proof.
     repeat split => //=; sfirstorder b:on use:ne_nf.
 Qed.
 
-
 Lemma algo_metric_bind n k p0 (A0 : PTm n) 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.
@@ -1164,7 +1258,7 @@ Proof.
   - split; last by sfirstorder use:hf_not_hne.
     move {hnfneu}.
     case : a h fb fa => //=.
-    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp.
+    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp.
       move => a0 a1 ha0 _ _ Γ A wt0 wt1.
       move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]].
       apply : CE_HRed; eauto using rtc_refl.
@@ -1195,7 +1289,7 @@ Proof.
       simpl in *.
       have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst.
       sfirstorder.
-    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
+    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp.
       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
       have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1)
         by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
@@ -1263,6 +1357,11 @@ Proof.
         eauto using Su_Eq.
       * move => > /algo_metric_join.
         hauto lq:on use:DJoin.bind_univ_noconf.
+      * move => > /algo_metric_join.
+        hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin.
+      * move => > /algo_metric_join.
+        admit.
+      * admit.
     + case : b => //=.
       * hauto lq:on use:T_AbsUniv_Imp.
       * hauto lq:on use:T_PairUniv_Imp.
@@ -1270,6 +1369,22 @@ Proof.
       * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj.
         subst.
         hauto l:on.
+      * move => > /algo_metric_join.
+        hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin.
+      * admit.
+      * admit.
+    + case : b => //=.
+      * qauto l:on use:T_AbsNat_Imp.
+      * qauto l:on use:T_PairNat_Imp.
+      * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf.
+      * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf.
+      * hauto l:on.
+      * admit.
+      * admit.
+    (* Zero *)
+    + admit.
+    (* Suc *)
+    + admit.
   - move : k a b h fb fa. abstract : hnfneu.
     move => k.
     move => + b.
@@ -1326,6 +1441,11 @@ Proof.
       hauto l:on use:Sub.hne_bind_noconf.
     (* NeuUniv: Impossible *)
     + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join.
+    + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join.
+    (* Zero *)
+    + admit.
+    (* Suc *)
+    + admit.
   - move {ih}.
     move /algo_metric_sym in h.
     qauto l:on use:coqeq_symmetric_mutual.
@@ -1349,6 +1469,7 @@ Proof.
       * move => a0 a1 i /algo_metric_join. clear => ? ? _. exfalso.
         hauto l:on use:REReds.var_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
+      * admit.
     + case : b => //=.
       * clear. move => i a a0 /algo_metric_join h _ ?. exfalso.
         hauto l:on use:REReds.hne_app_inv, REReds.var_inv.
@@ -1377,7 +1498,7 @@ Proof.
         move => [ih _].
         move : ih (ha0') (ha1'); repeat move/[apply].
         move => iha.
-        split. sauto lq:on.
+        split. qblast.
         exists (subst_PTm (scons a0 VarPTm) B2).
         split.
         apply : Su_Transitive; eauto.
@@ -1400,6 +1521,7 @@ Proof.
       * move => p p0 p1 p2 /algo_metric_join. clear => ? ? ?. exfalso.
         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
+      * admit.
     + case : b => //=.
       * move => i p h /algo_metric_join. clear => ? _ ?. exfalso.
         hauto l:on use:REReds.hne_proj_inv, REReds.var_inv.
@@ -1460,8 +1582,11 @@ Proof.
                move /regularity_sub0 in hSu21.
                sfirstorder use:bind_inst.
       * sfirstorder use:T_Bot_Imp.
+      * admit.
     + sfirstorder use:T_Bot_Imp.
-Qed.
+    (* ind ind case *)
+    + admit.
+Admitted.
 
 Lemma coqeq_sound : forall n Γ (a b : PTm n) A,
     Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b -> Γ ⊢ a ≡ b ∈ A.
@@ -1593,7 +1718,8 @@ Proof.
     inversion E; subst => /=.
     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-Qed.
+      (* Copy paste from algo_metric_case *)
+Admitted.
 
 Lemma CLE_HRedL n (a a' b : PTm n)  :
   HRed.R a a' ->
@@ -1629,7 +1755,7 @@ Proof.
     inversion E; subst => /=.
     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-Qed.
+Admitted.
 
 Lemma salgo_metric_sub n k (a b : PTm n) :
   salgo_metric k a b ->
@@ -1713,6 +1839,9 @@ Proof.
              by hauto l:on.
            eauto using ctx_eq_subst_one.
       * hauto lq:on use:salgo_metric_sub, Sub.bind_univ_noconf.
+      * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf.
+      * admit.
+      * admit.
     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
       * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
       * move => *. econstructor; eauto using rtc_refl.

From af0cac15e6853e6cd1f1c8fde43ee06d232902ea Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Wed, 26 Feb 2025 19:46:00 -0500
Subject: [PATCH 156/210] Prove some simple impossible cases

---
 theories/algorithmic.v | 61 ++++++++++++++++++++++++++++++++++++------
 theories/fp_red.v      | 23 ++++++++++++++++
 2 files changed, 76 insertions(+), 8 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index d7e226c..f714a7a 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -947,6 +947,16 @@ Proof.
   hauto lq:on use:Sub.bind_univ_noconf.
 Qed.
 
+Lemma lored_nsteps_suc_inv k n (a : PTm n) 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.
+  move : a E.
+  elim : u b /hu.
+  - hauto l:on.
+  - scrush ctrs:nsteps inv:LoRed.R.
+Qed.
+
 Lemma lored_nsteps_abs_inv k n (a : PTm (S n)) b :
   nsteps LoRed.R k (PAbs a) b -> exists b', nsteps LoRed.R k a b' /\ b = PAbs b'.
 Proof.
@@ -1211,6 +1221,14 @@ Proof.
     repeat split => //=; sfirstorder b:on use:ne_nf.
 Qed.
 
+Lemma algo_metric_suc n k (a0 a1 : PTm n) :
+  algo_metric k (PSuc a0) (PSuc a1) ->
+  exists j, j < k /\ algo_metric j a0 a1.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
+  exists (k - 1).
+
+
 Lemma algo_metric_bind n k p0 (A0 : PTm n) 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.
@@ -1360,8 +1378,9 @@ Proof.
       * move => > /algo_metric_join.
         hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin.
       * move => > /algo_metric_join.
-        admit.
-      * admit.
+        clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv.
+      * move => > /algo_metric_join. clear.
+        hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv.
     + case : b => //=.
       * hauto lq:on use:T_AbsUniv_Imp.
       * hauto lq:on use:T_PairUniv_Imp.
@@ -1371,20 +1390,46 @@ Proof.
         hauto l:on.
       * move => > /algo_metric_join.
         hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin.
-      * admit.
-      * admit.
+      * move => > /algo_metric_join.
+        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv.
+      * move => > /algo_metric_join.
+        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv.
     + case : b => //=.
       * qauto l:on use:T_AbsNat_Imp.
       * qauto l:on use:T_PairNat_Imp.
       * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf.
       * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf.
       * hauto l:on.
-      * admit.
-      * admit.
+      * move /algo_metric_join.
+        hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv.
     (* Zero *)
-    + admit.
+    + case : b => //=.
+      * hauto lq:on rew:off use:T_AbsZero_Imp.
+      * hauto lq: on use: T_PairZero_Imp.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv.
+      * hauto l:on.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv.
     (* Suc *)
-    + admit.
+    + case : b => //=.
+      * hauto lq:on rew:off use:T_AbsSuc_Imp.
+      * hauto lq:on use:T_PairSuc_Imp.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv.
+      * move => a0 a1 h _ _.
   - move : k a b h fb fa. abstract : hnfneu.
     move => k.
     move => + b.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 5ad793b..bffe1a7 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -2038,6 +2038,13 @@ Module EReds.
     - hauto lq:on rew:off ctrs:rtc, ERed.R inv:ERed.R, rtc.
   Qed.
 
+  Lemma zero_inv n (C : PTm n) :
+    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 :
     rtc ERed.R (PInd P a b c) C ->
     exists P0 a0 b0 c0,  rtc ERed.R P P0 /\ rtc ERed.R a a0 /\
@@ -2410,6 +2417,13 @@ 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) :
+    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 :
     rtc RERed.R (PSuc a) C ->
     exists b, rtc RERed.R a b /\ C = PSuc b.
@@ -2421,6 +2435,15 @@ 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 ->
+    C = PNat.
+  Proof.
+    move E : PNat => u hu. move : E.
+    elim : u C / hu=>//=.
+    hauto lq:on rew:off ctrs:rtc, RERed.R inv:RERed.R.
+  Qed.
+
 End REReds.
 
 Module LoRed.

From aaec92890230b23479b5271f9fc8db7781733a4f Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Thu, 27 Feb 2025 00:07:57 -0500
Subject: [PATCH 157/210] Make progress on the completeness proof

---
 theories/algorithmic.v | 35 ++++++++++++++++++++++++++++++++---
 1 file changed, 32 insertions(+), 3 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index f714a7a..1cff8ef 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1227,7 +1227,14 @@ Lemma algo_metric_suc n k (a0 a1 : PTm n) :
 Proof.
   move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
   exists (k - 1).
-
+  move /lored_nsteps_suc_inv : h0.
+  move => [a0'][ha0]?. subst.
+  move /lored_nsteps_suc_inv : h1.
+  move => [b0'][hb0]?. subst. simpl in *.
+  split; first by lia.
+  rewrite /algo_metric.
+  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  :
   algo_metric k (PBind p0 A0 B0) (PBind p1 A1 B1) ->
@@ -1430,6 +1437,13 @@ Proof.
       * move => > /algo_metric_join.
         hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv.
       * move => a0 a1 h _ _.
+        move /algo_metric_suc : h => [j [h4 h5]].
+        move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3].
+        move : ih h4 h5;do!move/[apply].
+        move => [ih _].
+        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 => + b.
@@ -1488,9 +1502,24 @@ Proof.
     + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join.
     + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join.
     (* Zero *)
-    + admit.
+    + case => //=.
+      * move => i /algo_metric_join. clear.
+        hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv.
+      * move => >/algo_metric_join. clear.
+        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.
     (* Suc *)
-    + admit.
+    + move => a0.
+      case => //=; move => >/algo_metric_join; clear.
+      * 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.
+      *
   - move {ih}.
     move /algo_metric_sym in h.
     qauto l:on use:coqeq_symmetric_mutual.

From e663a1a5963bbd29c5b65bf9bd4b12990577d8a4 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Thu, 27 Feb 2025 15:06:55 -0500
Subject: [PATCH 158/210] Finish most of the completeness proof

---
 theories/algorithmic.v | 168 +++++++++++++++++++++++++++++++++++------
 1 file changed, 147 insertions(+), 21 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 1cff8ef..9e6a04b 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -995,6 +995,37 @@ 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 :
+  nsteps LoRed.R k (PInd P0 a0 b0 c0) U ->
+  ishne a0 ->
+  exists iP ia ib ic P1 a1 b1 c1,
+    iP <= k /\ ia <= k /\ ib <= k /\ ic <= k /\
+      U = PInd P1 a1 b1 c1 /\
+      nsteps LoRed.R iP P0 P1 /\
+      nsteps LoRed.R ia a0 a1 /\
+      nsteps LoRed.R ib b0 b1 /\
+      nsteps LoRed.R ic c0 c1.
+Proof.
+  move E : (PInd P0 a0 b0 c0) => u hu.
+  move : P0 a0 b0 c0 E.
+  elim : k u U / hu.
+  - sauto lq:on.
+  - move => k t0 t1 t2 ht01 ht12 ih P0 a0 b0 c0 ? nea0. subst.
+    inversion ht01; subst => //=; spec_refl.
+    * move /(_ ltac:(done)) : ih.
+      move => [iP][ia][ib][ic].
+      exists (S iP), ia, ib, ic. sauto lq:on solve+:lia.
+    * move /(_ ltac:(sfirstorder use:lored_hne_preservation)) : ih.
+      move => [iP][ia][ib][ic].
+      exists iP, (S ia), ib, ic. sauto lq:on solve+:lia.
+    * move /(_ ltac:(done)) : ih.
+      move => [iP][ia][ib][ic].
+      exists iP, ia, (S ib), ic. sauto lq:on solve+:lia.
+    * move /(_ ltac:(done)) : ih.
+      move => [iP][ia][ib][ic].
+      exists iP, ia, ib, (S ic). sauto lq:on solve+:lia.
+Qed.
+
 Lemma lored_nsteps_app_inv k n (a0 b0 C : PTm n) :
     nsteps LoRed.R k (PApp a0 b0) C ->
     ishne a0 ->
@@ -1060,6 +1091,7 @@ Proof.
   lia.
 Qed.
 
+
 Lemma lored_nsteps_app_cong k n (a0 a1 b : PTm n) :
   nsteps LoRed.R k a0 a1 ->
   ishne a0 ->
@@ -1192,6 +1224,24 @@ 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 :
+  algo_metric k (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ algo_metric j P0 P1 /\ algo_metric j a0 a1 /\
+         algo_metric j b0 b1 /\ algo_metric j c0 c1.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 hne0 hne1.
+  move /lored_nsteps_ind_inv /(_ hne0) : h0.
+  move =>[iP][ia][ib][ic][P2][a2][b2][c2][?][?][?][?][?][?][?][?]?. subst.
+  move /lored_nsteps_ind_inv /(_ hne1) : h1.
+  move =>[iP0][ia0][ib0][ic0][P3][a3][b3][c3][?][?][?][?][?][?][?][?]?. subst.
+  move /EJoin.ind_inj : h4.
+  move => [?][?][?]?.
+  exists (k -1). simpl in *.
+  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) :
   algo_metric k (PApp a0 b0) (PApp a1 b1) ->
   ishne a0 ->
@@ -1519,7 +1569,8 @@ 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.
@@ -1532,21 +1583,22 @@ Proof.
     by firstorder.
 
     case : a h fb fa => //=.
-    + case : b => //=.
-      move => i j hi _ _.
-      * have ? : j = i by hauto lq:on use:algo_metric_join, DJoin.var_inj. subst.
-        move => Γ A B hA hB.
-        split. apply CE_VarCong.
-        exists (Γ i). hauto l:on use:Var_Inv, T_Var.
-      * move => p p0 f /algo_metric_join. clear => ? ? _. exfalso.
+    + case : b => //=; move => > /algo_metric_join.
+      * move /DJoin.var_inj => ?. subst. qauto l:on use:Var_Inv, T_Var,CE_VarCong.
+      * clear => ? ? _. exfalso.
         hauto l:on use:REReds.var_inv, REReds.hne_app_inv.
-      * move => a0 a1 i /algo_metric_join. clear => ? ? _. exfalso.
+      * clear => ? ? _. exfalso.
         hauto l:on use:REReds.var_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
-      * admit.
-    + case : b => //=.
-      * clear. move => i a a0 /algo_metric_join h _ ?. exfalso.
-        hauto l:on use:REReds.hne_app_inv, REReds.var_inv.
+      * clear => ? ? _. exfalso.
+        hauto q:on use:REReds.var_inv, REReds.hne_ind_inv.
+    + case : b => //=;
+                   lazymatch goal with
+                   | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac
+                   | _ => move => > /algo_metric_join
+                   end.
+      * clear => *. exfalso.
+        hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv.
       (* real case *)
       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
@@ -1592,10 +1644,11 @@ Proof.
         move /E_Symmetric in h01.
         move /regularity_sub0 : hSu20 => [i0].
         sfirstorder use:bind_inst.
-      * move => p p0 p1 p2 /algo_metric_join. clear => ? ? ?. exfalso.
+      * clear => ? ? ?. exfalso.
         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
       * sfirstorder use:T_Bot_Imp.
-      * admit.
+      * clear => ? ? ?. exfalso.
+        hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv.
     + case : b => //=.
       * move => i p h /algo_metric_join. clear => ? _ ?. exfalso.
         hauto l:on use:REReds.hne_proj_inv, REReds.var_inv.
@@ -1656,11 +1709,67 @@ Proof.
                move /regularity_sub0 in hSu21.
                sfirstorder use:bind_inst.
       * sfirstorder use:T_Bot_Imp.
-      * admit.
+      * 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 *)
-    + admit.
-Admitted.
+    + move => P a0 b0 c0.
+      case : b => //=;
+                   lazymatch goal with
+                   | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac
+                   | _ => move => > /algo_metric_join; clear => *; exfalso
+                   end.
+      * 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.
+        move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0.
+        move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1.
+        have {}iha : a0 ∼ a1 by qauto l:on.
+        have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia.
+        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.
+        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.
+        have haE : Γ ⊢ a0 ≡ a1 ∈ PNat
+          by hauto l:on use:coqeq_sound_mutual.
+        have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual.
+        have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)).
+        eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero.
+        have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P
+          by eauto using T_Conv_E.
+        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.
+        apply : T_Conv; eauto. apply : ctx_eq_subst_one; eauto with wt.
+        apply : Su_Eq; eauto.
+        subst T. apply : weakening_su; eauto.
+        eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto.
+        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.
+        have {}ihc : c0 ⇔ c1 by qauto l:on.
+        move => [:ih].
+        split. abstract : ih. move : h0 h1 ihP iha ihb ihc. clear. sauto lq:on.
+        have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual.
+        exists (subst_PTm (scons a0 VarPTm) P).
+        repeat split => //=; eauto with wt.
+        apply : Su_Transitive.
+        apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt.
+        apply : Su_Eq. apply E_Refl. by apply T_Nat'.
+        apply : Su_Eq. apply hPE. by eauto.
+        move : hEInd. clear. hauto l:on use:regularity.
+Qed.
 
 Lemma coqeq_sound : forall n Γ (a b : PTm n) A,
     Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b -> Γ ⊢ a ≡ b ∈ A.
@@ -1792,8 +1901,16 @@ Proof.
     inversion E; subst => /=.
     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-      (* Copy paste from algo_metric_case *)
-Admitted.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + sfirstorder use:ne_hne.
+Qed.
 
 Lemma CLE_HRedL n (a a' b : PTm n)  :
   HRed.R a a' ->
@@ -1829,7 +1946,16 @@ Proof.
     inversion E; subst => /=.
     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-Admitted.
+  - inversion 1 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + sfirstorder use:ne_hne.
+Qed.
 
 Lemma salgo_metric_sub n k (a b : PTm n) :
   salgo_metric k a b ->

From 4509a64bf5004597d484ad617b86939a244ffd56 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Thu, 27 Feb 2025 15:30:55 -0500
Subject: [PATCH 159/210] Finish the soundness and completeness proof with nat

---
 theories/algorithmic.v | 39 +++++++++++++++++++++++++++++++++++++--
 1 file changed, 37 insertions(+), 2 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 9e6a04b..096ec94 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1815,6 +1815,9 @@ Inductive CoqLEq {n} : PTm n -> PTm n -> Prop :=
   (* ---------------------------- *)
   PBind PSig A0 B0 ⋖ PBind PSig A1 B1
 
+| CLE_NatCong :
+  PNat ⋖ PNat
+
 | CLE_NeuNeu a0 a1 :
   a0 ∼ a1 ->
   a0 ⋖ a1
@@ -1877,6 +1880,7 @@ Proof.
     apply : ihB; by eauto using ctx_eq_subst_one.
     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.
     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.
@@ -2040,23 +2044,54 @@ Proof.
            eauto using ctx_eq_subst_one.
       * hauto lq:on use:salgo_metric_sub, Sub.bind_univ_noconf.
       * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf.
-      * admit.
-      * admit.
+      * move => _ > _ /salgo_metric_sub.
+        hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
       * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
       * move => *. econstructor; eauto using rtc_refl.
         hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong.
+      * hauto lq:on rew:off use:REReds.univ_inv, REReds.nat_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
   (* Both cases are impossible *)
+    + case : b fb => //=.
+      * hauto lq:on use:T_AbsNat_Imp.
+      * hauto lq:on use:T_PairNat_Imp.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.bind_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto l:on.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
+    + move => ? ? Γ i /Zero_Inv.
+      clear.
+      move /synsub_to_usub => [_ [_ ]].
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
+    + move => ? _ _ Γ i /Suc_Inv => [[_]].
+      move /synsub_to_usub.
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
   - have {}h : DJoin.R a b by
        hauto lq:on use:salgo_metric_algo_metric, algo_metric_join.
     case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
     + hauto lq:on use:DJoin.hne_bind_noconf.
     + hauto lq:on use:DJoin.hne_univ_noconf.
+    + hauto lq:on use:DJoin.hne_nat_noconf.
+    + move => _ _ Γ i _.
+      move /Zero_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
+    + move => p _ _ Γ i _ /Suc_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
   - have {}h : DJoin.R b a by
       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric.
     case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
     + hauto lq:on use:DJoin.hne_bind_noconf.
     + hauto lq:on use:DJoin.hne_univ_noconf.
+    + hauto lq:on use:DJoin.hne_nat_noconf.
+    + move => _ _ Γ i.
+      move /Zero_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
+    + move => p _ _ Γ i /Suc_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
   - move => Γ i ha hb.
     econstructor; eauto using rtc_refl.
     apply CLE_NeuNeu. move {ih}.

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 160/210] 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 161/210] 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 162/210] 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 163/210] 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 164/210] 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 165/210] 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 166/210] 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 167/210] 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 168/210] 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 169/210] 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 170/210] 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 171/210] 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 172/210] 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 173/210] 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 174/210] 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 175/210] 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 176/210] 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 177/210] 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 178/210] 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 179/210] 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 180/210] 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 181/210] 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 182/210] 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 183/210] 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.

From 96b31397265910c301fe6fd10adeb6578be2ee3e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Mar 2025 15:39:27 -0500
Subject: [PATCH 184/210] Prove termination

---
 theories/algorithmic.v        |   6 +
 theories/executable.v         |   5 +
 theories/executable_correct.v |  13 --
 theories/termination.v        | 258 +++++++++++++++++++++++++++++++++-
 4 files changed, 266 insertions(+), 16 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 8a9146a..ab8fe97 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1050,6 +1050,12 @@ Proof.
 Qed.
 
 
+Lemma hreds_nf_refl a b  :
+  HRed.nf a ->
+  rtc HRed.R a b ->
+  a = b.
+Proof. inversion 2; sfirstorder. Qed.
+
 Lemma lored_nsteps_app_cong k (a0 a1 b : PTm) :
   nsteps LoRed.R k a0 a1 ->
   ishne a0 ->
diff --git a/theories/executable.v b/theories/executable.v
index 0586a41..156e691 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -486,3 +486,8 @@ Lemma check_equal_conf a b nfa nfb nfab :
 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.
+
+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.
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 0720d03..fdfee5a 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -4,12 +4,6 @@ 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 coqeqr_no_hred a b : a ∼ b -> HRed.nf a /\ HRed.nf b.
 Proof. induction 1; sauto lq:on unfold:HRed.nf. Qed.
 
@@ -18,7 +12,6 @@ Proof. induction 1;
          qauto inv:HRed.R use:coqeqr_no_hred, hne_no_hred unfold:HRed.nf.
 Qed.
 
-
 Lemma coqeq_neuneu u0 u1 :
   ishne u0 -> ishne u1 -> u0 ↔ u1 -> u0 ∼ u1.
 Proof.
@@ -88,12 +81,6 @@ 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.
-
 Ltac ce_solv := move => *; simp ce_prop in *; hauto qb:on rew:off inv:CoqEq, CoqEq_Neu.
 
 Lemma check_equal_complete :
diff --git a/theories/termination.v b/theories/termination.v
index 61f1e1d..53474b6 100644
--- a/theories/termination.v
+++ b/theories/termination.v
@@ -1,14 +1,266 @@
 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 (..)).
+From stdpp Require Import relations (nsteps (..), rtc(..)).
+Import Psatz.
+
+Lemma tm_conf_sym a b : tm_conf a b = tm_conf b a.
+Proof. case : a; case : b => //=. Qed.
+
+#[export]Hint Constructors algo_dom algo_dom_r : adom.
+
+Lemma algo_dom_r_inv a b :
+  algo_dom_r a b -> exists a0 b0, algo_dom a0 b0 /\ rtc HRed.R a a0 /\ rtc HRed.R b b0.
+Proof.
+  induction 1; hauto lq:on ctrs:rtc.
+Qed.
+
+Lemma A_HRedsL a a' b :
+  rtc HRed.R a a' ->
+  algo_dom_r a' b ->
+  algo_dom_r a b.
+  induction 1; sfirstorder use:A_HRedL.
+Qed.
+
+Lemma A_HRedsR a b b' :
+  HRed.nf a ->
+  rtc HRed.R b b' ->
+  algo_dom a b' ->
+  algo_dom_r a b.
+Proof. induction 2; sauto. Qed.
+
+Lemma algo_dom_sym :
+  (forall a b (h : algo_dom a b), algo_dom b a) /\
+    (forall a b (h : algo_dom_r a b), algo_dom_r b a).
+Proof.
+  apply algo_dom_mutual; try qauto use:tm_conf_sym db:adom.
+  move => a a' b hr h ih.
+  move /algo_dom_r_inv : ih => [a0][b0][ih0][ih1]ih2.
+  have nfa0 : HRed.nf a0 by sfirstorder use:algo_dom_no_hred.
+  sauto use:A_HRedsL, A_HRedsR.
+Qed.
 
 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_sym k (a b : PTm) :
+  term_metric k a b -> term_metric k b a.
+Proof. hauto lq:on unfold:term_metric solve+:lia. Qed.
+
+Lemma term_metric_case k (a b : PTm) :
+  term_metric k a b ->
+  (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ term_metric k' a' b /\ k' < k.
+Proof.
+  move=>[i][j][va][vb][h0] [h1][h2][h3]h4.
+  case : a h0 => //=; try firstorder.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.AppAbs unfold:term_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:term_metric solve+:lia.
+    + sfirstorder qb:on use:ne_hne.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.ProjPair unfold:term_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:term_metric solve+:lia.
+  - inversion h0 as [|A B C D E F]; subst.
+    hauto qb:on use:ne_hne.
+    inversion E; subst => /=.
+    + hauto lq:on use:HRed.IndZero unfold:term_metric solve+:lia.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:term_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:term_metric solve+:lia.
+    + sfirstorder use:ne_hne.
+    + sfirstorder use:ne_hne.
+Qed.
+
+Lemma A_Conf' a b :
+  ishf a \/ ishne a ->
+  ishf b \/ ishne b ->
+  tm_conf a b ->
+  algo_dom_r a b.
+Proof.
+  move => ha hb.
+  have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
+  have {}hb : HRed.nf b by sfirstorder use:hf_no_hred, hne_no_hred.
+  move => ?.
+  apply A_NfNf.
+  by apply A_Conf.
+Qed.
+
+Lemma term_metric_abs : forall k a b,
+    term_metric k (PAbs a) (PAbs b) ->
+    exists k', k' < k /\ term_metric k' a b.
+Proof.
+  move => k a b [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_abs_inv in hva, hvb.
+  move : hva => [a'][hva]?. subst.
+  move : hvb => [b'][hvb]?. subst.
+  simpl in *. exists (k - 1).
+  hauto lq:on unfold:term_metric solve+:lia.
+Qed.
+
+Lemma term_metric_pair : forall k a0 b0 a1 b1,
+    term_metric k (PPair a0 b0) (PPair a1 b1) ->
+    exists k', k' < k /\ term_metric k' a0 a1 /\ term_metric k' b0 b1.
+Proof.
+  move => k a0 b0 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_pair_inv in hva, hvb.
+  move : hva => [?][?][?][?][?][?][?][?]?.
+  move : hvb => [?][?][?][?][?][?][?][?]?. subst.
+  simpl in *. exists (k - 1).
+  hauto lqb:on solve+:lia.
+Qed.
+
+Lemma term_metric_bind : forall k p0 a0 b0 p1 a1 b1,
+    term_metric k (PBind p0 a0 b0) (PBind p1 a1 b1) ->
+    exists k', k' < k /\ term_metric k' a0 a1 /\ term_metric k' b0 b1.
+Proof.
+  move => k p0 a0 b0 p1 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_bind_inv in hva, hvb.
+  move : hva => [?][?][?][?][?][?][?][?]?.
+  move : hvb => [?][?][?][?][?][?][?][?]?. subst.
+  simpl in *. exists (k - 1).
+  hauto lqb:on solve+:lia.
+Qed.
+
+Lemma term_metric_suc : forall k a b,
+    term_metric k (PSuc a) (PSuc b) ->
+    exists k', k' < k /\ term_metric k' a b.
+Proof.
+  move => k a b [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_suc_inv in hva, hvb.
+  move : hva => [a'][hva]?. subst.
+  move : hvb => [b'][hvb]?. subst.
+  simpl in *. exists (k - 1).
+  hauto lq:on unfold:term_metric solve+:lia.
+Qed.
+
+Lemma term_metric_abs_neu k (a0 : PTm) u :
+  term_metric k (PAbs a0) u ->
+  ishne u ->
+  exists j, j < k /\ term_metric j a0 (PApp (ren_PTm shift u) (VarPTm var_zero)).
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 neu.
+  have neva : ne vb by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
+  move /lored_nsteps_abs_inv : h0 => [a1][h01]?. subst.
+  exists (k - 1).
+  simpl in *. split. lia.
+  exists i,j,a1,(PApp (ren_PTm shift vb) (VarPTm var_zero)).
+  repeat split => //=.
+  apply lored_nsteps_app_cong.
+  by apply lored_nsteps_renaming.
+  by rewrite ishne_ren.
+  rewrite Bool.andb_true_r.
+  sfirstorder use:ne_nf_ren.
+  rewrite size_PTm_ren. lia.
+Qed.
+
+Lemma term_metric_pair_neu k (a0 b0 : PTm) u :
+  term_metric k (PPair a0 b0) u ->
+  ishne u ->
+  exists j, j < k /\ term_metric j (PProj PL u) a0 /\ term_metric j (PProj PR u) b0.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 neu.
+  have neva : ne vb by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
+  move /lored_nsteps_pair_inv : h0 => [i0][j0][a1][b1][?][?][?][?]?. subst.
+  exists (k-1). sauto qb:on use:lored_nsteps_proj_cong unfold:term_metric solve+:lia.
+Qed.
+
+Lemma term_metric_app k (a0 b0 a1 b1 : PTm) :
+  term_metric k (PApp a0 b0) (PApp a1 b1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ term_metric j a0 a1 /\ term_metric j b0 b1.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4.
+  move => hne0 hne1.
+  move : lored_nsteps_app_inv h0 (hne0);repeat move/[apply].
+  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
+  move : lored_nsteps_app_inv h1 (hne1);repeat move/[apply].
+  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+  simpl in *. exists (k - 1). hauto lqb:on use:lored_nsteps_app_cong, ne_nf unfold:term_metric solve+:lia.
+Qed.
+
+Lemma term_metric_proj k p0 p1 (a0 a1 : PTm) :
+  term_metric k (PProj p0 a0) (PProj p1 a1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ term_metric j a0 a1.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 hne0 hne1.
+  move : lored_nsteps_proj_inv h0 (hne0);repeat move/[apply].
+  move => [i0][a2][hi][?]ha02. subst.
+  move : lored_nsteps_proj_inv h1 (hne1);repeat move/[apply].
+  move => [i1][a3][hj][?]ha13. subst.
+  exists (k- 1).  hauto q:on use:ne_nf solve+:lia.
+Qed.
+
+Lemma term_metric_ind k P0 (a0 : PTm ) b0 c0 P1 a1 b1 c1 :
+  term_metric k (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ term_metric j P0 P1 /\ term_metric j a0 a1 /\
+         term_metric j b0 b1 /\ term_metric j c0 c1.
+Proof.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 hne0 hne1.
+  move /lored_nsteps_ind_inv /(_ hne0) : h0.
+  move =>[iP][ia][ib][ic][P2][a2][b2][c2][?][?][?][?][?][?][?][?]?. subst.
+  move /lored_nsteps_ind_inv /(_ hne1) : h1.
+  move =>[iP0][ia0][ib0][ic0][P3][a3][b3][c3][?][?][?][?][?][?][?][?]?. subst.
+  exists (k -1). simpl in *.
+  hauto lq:on rew:off use:ne_nf b:on solve+:lia.
+Qed.
+
+
+Lemma algo_dom_r_algo_dom : forall a b, HRed.nf a -> HRed.nf b -> algo_dom_r a b -> algo_dom a b.
+Proof. hauto l:on use:algo_dom_r_inv, hreds_nf_refl. Qed.
+
 Lemma term_metric_algo_dom : forall k a b, term_metric k a b -> algo_dom_r a b.
 Proof.
+  move => [:hneL].
   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.
+  case /term_metric_case : (h); cycle 1.
+  move => [k'][a'][h0][h1]h2.
+  by apply : A_HRedL; eauto.
+  case /term_metric_sym /term_metric_case : (h); cycle 1.
+  move => [k'][b'][hb][/term_metric_sym h0]h1.
+  move => ha. have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
+  by apply : A_HRedR; eauto.
+  move => /[swap].
+  case => hfa; case => hfb.
+  - move : hfa hfb h.
+    case : a; case : b => //=; eauto 5 using A_Conf' with adom.
+    + hauto lq:on use:term_metric_abs db:adom.
+    + hauto lq:on use:term_metric_pair db:adom.
+    + hauto lq:on use:term_metric_bind db:adom.
+    + hauto lq:on use:term_metric_suc db:adom.
+  - abstract : hneL n ih a b h hfa hfb.
+    case : a hfa h => //=.
+    + hauto lq:on use:term_metric_abs_neu db:adom.
+    + scrush use:term_metric_pair_neu db:adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+  - hauto q:on use:algo_dom_sym, term_metric_sym.
+  - move {hneL}.
+    case : b hfa hfb h => //=; case a => //=; eauto 5 using A_Conf' with adom.
+    + move => a0 b0 a1 b1 nfa0 nfa1.
+      move /term_metric_app /(_ nfa0  nfa1) => [j][hj][ha]hb.
+      apply A_NfNf. apply A_AppCong => //; eauto.
+      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
+      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
+      eauto using algo_dom_r_algo_dom.
+    + move => p0 A0 p1 A1 neA0 neA1.
+      have {}nfa0 : HRed.nf A0 by sfirstorder use:hne_no_hred.
+      have {}nfb0 : HRed.nf A1 by sfirstorder use:hne_no_hred.
+      hauto lq:on use:term_metric_proj, algo_dom_r_algo_dom db:adom.
+    + move => P0 a0 b0 c0 P1 a1 b1 c1 nea0 nea1.
+      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
+      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
+      hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
+Qed.

From 6f154cc9c61961a896a260b2dede10b44bf7890b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Mar 2025 16:20:32 -0500
Subject: [PATCH 185/210] Add stub for check_sub

---
 theories/executable.v  | 22 ++++++++++++++++++----
 theories/termination.v | 16 ++++++++++++++++
 2 files changed, 34 insertions(+), 4 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index 156e691..a6fa438 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -212,6 +212,11 @@ with algo_dom_r : PTm  -> PTm -> Prop :=
   (* ----------------------- *)
   algo_dom_r a b.
 
+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 hf_no_hred (a b : PTm) :
   ishf a ->
   HRed.R a b ->
@@ -351,7 +356,6 @@ Ltac solve_check_equal :=
         | inr b' := check_equal_r a (proj1_sig b') _;
                  | inl h'' := check_equal a b _.
 
-
 Next Obligation.
   intros.
   inversion h; subst => //=.
@@ -487,7 +491,17 @@ 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.
 
-Scheme algo_ind := Induction for algo_dom Sort Prop
-  with algor_ind := Induction for algo_dom_r Sort Prop.
+Obligation Tactic := solve [check_equal_triv | sfirstorder].
 
-Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
+Program Fixpoint check_sub (q : bool) (a b : PTm) (h : algo_dom a b) {struct h} :=
+  match a, b with
+  | PBind PPi A0 B0, PBind PPi A1 B1 =>
+      check_sub_r (negb q) A0 A1 _ && check_sub_r q B0 B1 _
+  | PBind PSig A0 B0, PBind PSig A1 B1 =>
+      check_sub_r q A0 A1 _ && check_sub_r q B0 B1 _
+  | PUniv i, PUniv j =>
+      if q then PeanoNat.Nat.leb i j else PeanoNat.Nat.leb j i
+  | PNat, PNat => true
+  | _ ,_ => ishne a && ishne b && check_equal a b h
+  end
+with check_sub_r (q : bool) (a b : PTm) (h : algo_dom_r a b) {struct h} := false.
diff --git a/theories/termination.v b/theories/termination.v
index 53474b6..0f8314d 100644
--- a/theories/termination.v
+++ b/theories/termination.v
@@ -264,3 +264,19 @@ Proof.
       have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
       hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
 Qed.
+
+Lemma sn_term_metric (a b : PTm) : SN a -> SN b -> exists k, term_metric k a b.
+Proof.
+  move /LoReds.FromSN => [va [ha0 ha1]].
+  move /LoReds.FromSN => [vb [hb0 hb1]].
+  eapply relations.rtc_nsteps in ha0.
+  eapply relations.rtc_nsteps in hb0.
+  hauto lq:on unfold:term_metric solve+:lia.
+Qed.
+
+Lemma sn_algo_dom a b : SN a -> SN b -> algo_dom_r a b.
+Proof.
+  move : sn_term_metric; repeat move/[apply].
+  move => [k]+.
+  eauto using term_metric_algo_dom.
+Qed.

From fe52d78ec5b96c6a5db5919b120c35db3091c9db Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Mar 2025 19:20:54 -0500
Subject: [PATCH 186/210] Start the soundness proof for check_sub

---
 theories/executable.v         | 101 +++++++++++++++++++++++++++++++++-
 theories/executable_correct.v |  23 ++++++++
 2 files changed, 122 insertions(+), 2 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index a6fa438..4ca9c5b 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -491,7 +491,7 @@ 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.
 
-Obligation Tactic := solve [check_equal_triv | sfirstorder].
+Obligation Tactic := try solve [check_equal_triv | sfirstorder].
 
 Program Fixpoint check_sub (q : bool) (a b : PTm) (h : algo_dom a b) {struct h} :=
   match a, b with
@@ -504,4 +504,101 @@ Program Fixpoint check_sub (q : bool) (a b : PTm) (h : algo_dom a b) {struct h}
   | PNat, PNat => true
   | _ ,_ => ishne a && ishne b && check_equal a b h
   end
-with check_sub_r (q : bool) (a b : PTm) (h : algo_dom_r a b) {struct h} := false.
+with check_sub_r (q : bool) (a b : PTm) (h : algo_dom_r a b) {struct h} :=
+  match fancy_hred a with
+  | inr a' => check_sub_r q (proj1_sig a') b _
+  | inl ha' => match fancy_hred b with
+            | inr b' => check_sub_r q a (proj1_sig b') _
+            | inl hb' => check_sub q a b _
+            end
+  end.
+Next Obligation.
+  simpl. intros. clear Heq_anonymous.  destruct a' as [a' ha']. simpl.
+  inversion h; subst => //=.
+  exfalso. sfirstorder use:algo_dom_no_hred.
+  assert (a' = a'0) by eauto using hred_deter. by subst.
+  exfalso. sfirstorder.
+Defined.
+
+Next Obligation.
+  simpl. intros. clear Heq_anonymous Heq_anonymous0.
+  destruct b' as [b' hb']. simpl.
+  inversion h; subst.
+  - exfalso.
+    sfirstorder use:algo_dom_no_hred.
+  - exfalso.
+    sfirstorder.
+  - assert (b' = b'0) by eauto using hred_deter. by subst.
+Defined.
+
+(* Need to avoid ssreflect tactics since they generate terms that make the termination checker upset *)
+Next Obligation.
+  move => _ /= a b hdom ha _ hb _.
+  inversion hdom; subst.
+  - assumption.
+  - exfalso; sfirstorder.
+  - exfalso; sfirstorder.
+Defined.
+
+Lemma check_sub_pi_pi q A0 B0  A1 B1 h0 h1 :
+  check_sub q (PBind PPi A0 B0) (PBind PPi A1 B1) (A_BindCong PPi PPi A0 A1 B0 B1 h0 h1) =
+    check_sub_r (~~ q) A0 A1 h0 && check_sub_r q B0 B1 h1.
+Proof. reflexivity. Qed.
+
+Lemma check_sub_sig_sig q A0 B0  A1 B1 h0 h1 :
+  check_sub q (PBind PSig A0 B0) (PBind PSig A1 B1) (A_BindCong PSig PSig A0 A1 B0 B1 h0 h1) =
+    check_sub_r q A0 A1 h0 && check_sub_r q B0 B1 h1.
+  reflexivity. Qed.
+
+Lemma check_sub_univ_univ i j :
+  check_sub true (PUniv i) (PUniv j) (A_UnivCong i j) = PeanoNat.Nat.leb i j.
+Proof. reflexivity. Qed.
+
+Lemma check_sub_univ_univ' i j :
+  check_sub false (PUniv i) (PUniv j) (A_UnivCong i j) = PeanoNat.Nat.leb j i.
+  reflexivity. Qed.
+
+Lemma check_sub_nfnf q a b dom : check_sub_r q a b (A_NfNf a b dom) = check_sub q 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.
+  destruct (fancy_hred b)=>//=.
+  destruct (fancy_hred a) =>//=.
+  destruct s as [a' ha']. simpl.
+  hauto l:on use:hred_complete.
+  destruct s as [b' hb']. simpl.
+  hauto l:on use:hred_complete.
+Qed.
+
+Lemma check_sub_hredl q a b a' ha doma :
+  check_sub_r q a b (A_HRedL a a' b ha doma) = check_sub_r q a' b doma.
+Proof.
+  simpl.
+  destruct (fancy_hred a).
+  - hauto q:on unfold:HRed.nf.
+  - destruct s as [x ?].
+    have ? : x = a' by eauto using hred_deter. subst.
+    simpl.
+    f_equal.
+    apply PropExtensionality.proof_irrelevance.
+Qed.
+
+Lemma check_sub_hredr q a b b' hu r a0 :
+  check_sub_r q a b (A_HRedR a b b' hu r a0) =
+    check_sub_r q a b' a0.
+Proof.
+  simpl.
+  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_sub_hredl check_sub_hredr check_sub_nfnf check_sub_univ_univ' check_sub_univ_univ check_sub_pi_pi check_sub_sig_sig : ce_prop.
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index fdfee5a..3cc6233 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -164,3 +164,26 @@ Proof.
       sfirstorder unfold:HRed.nf.
     + sauto lq:on use:hred_deter.
 Qed.
+
+Ltac simp_sub := with_strategy opaque [check_equal] simpl.
+
+Lemma check_sub_sound :
+  (forall a b (h : algo_dom a b), forall q, check_sub q a b h -> if q then a ⋖ b else b ⋖ a) /\
+    (forall a b (h : algo_dom_r a b), forall q, check_sub_r q a b h -> if q then a ≪ b else b ≪ a).
+Proof.
+  apply algo_dom_mutual; try done.
+  - move => a [] //=; hauto qb:on.
+  - move => a0 a1 []//=; hauto qb:on.
+  - simpl. move => i j [];
+    sauto lq:on use:Reflect.Nat_leb_le.
+  - admit.
+  - hauto l:on.
+  - move => i j q h.
+    have {}h : nat_eqdec i j by sfirstorder.
+    case : nat_eqdec h => //=; sauto lq:on.
+  - simp_sub.
+    move => p0 p1 u0 u1 i i0 dom ihdom q.
+    move /andP => [/andP [h00 h01] h1].
+    best use:check_sub_
+
+best b:on use:check_equal_sound.

From 437c97455e7d55255349d02b26071e831b1c2be3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 6 Mar 2025 20:42:08 -0500
Subject: [PATCH 187/210] Finish the soundness and completeness proof of the
 subtyping algorithm

---
 theories/executable.v         |  56 +++++++++++++++
 theories/executable_correct.v | 125 ++++++++++++++++++++++++++++++++--
 theories/termination.v        |   8 +--
 3 files changed, 179 insertions(+), 10 deletions(-)

diff --git a/theories/executable.v b/theories/executable.v
index 4ca9c5b..44557f9 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -217,6 +217,62 @@ Scheme algo_ind := Induction for algo_dom Sort Prop
 
 Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
 
+
+(* Inductive salgo_dom : PTm -> PTm -> Prop := *)
+(* | S_UnivCong i j : *)
+(*   (* -------------------------- *) *)
+(*   salgo_dom (PUniv i) (PUniv j) *)
+
+(* | S_PiCong A0 A1 B0 B1 : *)
+(*   salgo_dom_r A1 A0 -> *)
+(*   salgo_dom_r B0 B1 -> *)
+(*   (* ---------------------------- *) *)
+(*   salgo_dom (PBind PPi A0 B0) (PBind PPi A1 B1) *)
+
+(* | S_SigCong A0 A1 B0 B1 : *)
+(*   salgo_dom_r A0 A1 -> *)
+(*   salgo_dom_r B0 B1 -> *)
+(*   (* ---------------------------- *) *)
+(*   salgo_dom (PBind PSig A0 B0) (PBind PSig A1 B1) *)
+
+(* | S_NatCong : *)
+(*   salgo_dom PNat PNat *)
+
+(* | S_NeuNeu a b : *)
+(*   ishne a -> *)
+(*   ishne b -> *)
+(*   algo_dom a b -> *)
+(*   (* ------------------- *) *)
+(*   salgo_dom *)
+
+(* | S_Conf a b : *)
+(*   HRed.nf a -> *)
+(*   HRed.nf b -> *)
+(*   tm_conf a b -> *)
+(*   salgo_dom a b *)
+
+(* with salgo_dom_r : PTm  -> PTm -> Prop := *)
+(* | S_NfNf a b : *)
+(*   salgo_dom a b -> *)
+(*   salgo_dom_r a b *)
+
+(* | S_HRedL a a' b  : *)
+(*   HRed.R a a' -> *)
+(*   salgo_dom_r a' b -> *)
+(*   (* ----------------------- *) *)
+(*   salgo_dom_r a b *)
+
+(* | S_HRedR a b b'  : *)
+(*   HRed.nf a -> *)
+(*   HRed.R b b' -> *)
+(*   salgo_dom_r a b' -> *)
+(*   (* ----------------------- *) *)
+(*   salgo_dom_r a b. *)
+
+(* Scheme salgo_ind := Induction for salgo_dom Sort Prop *)
+(*   with algor_ind := Induction for salgo_dom_r Sort Prop. *)
+
+
 Lemma hf_no_hred (a b : PTm) :
   ishf a ->
   HRed.R a b ->
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 3cc6233..6ccee46 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -12,6 +12,11 @@ Proof. induction 1;
          qauto inv:HRed.R use:coqeqr_no_hred, hne_no_hred unfold:HRed.nf.
 Qed.
 
+Lemma coqleq_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, coqeqr_no_hred unfold:HRed.nf.
+Qed.
+
 Lemma coqeq_neuneu u0 u1 :
   ishne u0 -> ishne u1 -> u0 ↔ u1 -> u0 ∼ u1.
 Proof.
@@ -165,8 +170,9 @@ Proof.
     + sauto lq:on use:hred_deter.
 Qed.
 
-Ltac simp_sub := with_strategy opaque [check_equal] simpl.
+Ltac simp_sub := with_strategy opaque [check_equal] simpl in *.
 
+(* Reusing algo_dom results in an inefficient proof, but i'll brute force it so i can get the result more quickly *)
 Lemma check_sub_sound :
   (forall a b (h : algo_dom a b), forall q, check_sub q a b h -> if q then a ⋖ b else b ⋖ a) /\
     (forall a b (h : algo_dom_r a b), forall q, check_sub_r q a b h -> if q then a ≪ b else b ≪ a).
@@ -176,14 +182,121 @@ Proof.
   - move => a0 a1 []//=; hauto qb:on.
   - simpl. move => i j [];
     sauto lq:on use:Reflect.Nat_leb_le.
-  - admit.
+  - move => p0 p1 A0 A1 B0 B1 a iha b ihb q.
+    case : p0; case : p1; try done; simp ce_prop.
+    sauto lqb:on.
+    sauto lqb:on.
   - hauto l:on.
   - move => i j q h.
     have {}h : nat_eqdec i j by sfirstorder.
     case : nat_eqdec h => //=; sauto lq:on.
   - simp_sub.
-    move => p0 p1 u0 u1 i i0 dom ihdom q.
-    move /andP => [/andP [h00 h01] h1].
-    best use:check_sub_
+    sauto lqb:on use:coqeq_symmetric_mutual, coqeq_neuneu, check_equal_sound.
+  - simp_sub.
+    sauto lqb:on use:coqeq_symmetric_mutual, coqeq_neuneu, check_equal_sound.
+  - simp_sub.
+    sauto lqb:on use:coqeq_symmetric_mutual, coqeq_neuneu, check_equal_sound.
+  - move => a b n n0 i q h.
+    exfalso.
+    destruct a, b; try done; simp_sub; hauto lb:on use:check_equal_conf.
+  - move => a b doma ihdom q.
+    simp ce_prop. sauto lq:on.
+  - move => a a' b hr doma ihdom q.
+    simp ce_prop. move : ihdom => /[apply]. move {doma}.
+    sauto lq:on.
+  - move => a b b' n r dom ihdom q.
+    simp ce_prop.
+    move : ihdom => /[apply].
+    move {dom}.
+    sauto lq:on rew:off.
+Qed.
 
-best b:on use:check_equal_sound.
+Lemma check_sub_complete :
+  (forall a b (h : algo_dom a b), forall q, check_sub q a b h = false -> if q then ~ a ⋖ b else ~ b ⋖ a) /\
+    (forall a b (h : algo_dom_r a b), forall q, check_sub_r q a b h = false -> if q then ~ a ≪ b else ~ b ≪ a).
+Proof.
+  apply algo_dom_mutual; try first [done | hauto depth:4 lq:on inv:CoqLEq, CoqEq_Neu].
+  - move => i j q /=.
+    sauto lq:on rew:off use:PeanoNat.Nat.leb_le.
+  - move => p0 p1 A0 A1 B0 B1 a iha b ihb [].
+    + move => + h. inversion h; subst; simp ce_prop.
+      * move /nandP.
+        case.
+        by move => /negbTE {}/iha.
+        by move => /negbTE {}/ihb.
+      * move /nandP.
+        case.
+        by move => /negbTE {}/iha.
+        by move => /negbTE {}/ihb.
+      * inversion H.
+    + move => + h. inversion h; subst; simp ce_prop.
+      * move /nandP.
+        case.
+        by move => /negbTE {}/iha.
+        by move => /negbTE {}/ihb.
+      * move /nandP.
+        case.
+        by move => /negbTE {}/iha.
+        by move => /negbTE {}/ihb.
+      * inversion H.
+  - simp_sub.
+    sauto lq:on use:check_equal_complete.
+  - simp_sub.
+    move => p0 p1 u0 u1 i i0 a iha q.
+    move /nandP.
+    case.
+    move /nandP.
+    case => //.
+    by move /negP.
+    by move /negP.
+    move /negP.
+    move => h. eapply check_equal_complete in h.
+    sauto use:coqeq_symmetric_mutual inv:CoqLEq lq:on.
+  - move => u0 u1 a0 a1 i i0 a hdom ihdom hdom' ihdom'.
+    simp_sub.
+    move /nandP.
+    case.
+    move/nandP.
+    case.
+    by move/negP.
+    by move/negP.
+    move /negP.
+    move => h. eapply check_equal_complete in h.
+    sauto use:coqeq_symmetric_mutual inv:CoqLEq lq:on.
+  - move => P0 P1 u0 u1 b0 b1 c0 c1 i i0 dom ihdom dom' ihdom' dom'' ihdom'' dom''' ihdom''' q.
+    move /nandP.
+    case.
+    move /nandP.
+    case => //=.
+    by move/negP.
+    by move/negP.
+    move /negP => h. eapply check_equal_complete in h.
+    sauto use:coqeq_symmetric_mutual inv:CoqLEq lq:on.
+  - move => a b h ih q. simp ce_prop => {}/ih.
+    case : q => h0;
+    inversion 1; subst; hauto l:on use:algo_dom_no_hred, hreds_nf_refl.
+  - move => a a' b r dom ihdom q.
+    simp ce_prop => {}/ihdom.
+    case : q => h0.
+    inversion 1; subst.
+    inversion H0; subst. sfirstorder use:coqleq_no_hred.
+    have ? : a' = y by eauto using hred_deter. subst.
+    sauto lq:on.
+    inversion 1; subst.
+    inversion H1; subst. sfirstorder use:coqleq_no_hred.
+    have ? : a' = y by eauto using hred_deter. subst.
+    sauto lq:on.
+  - move => a b b' n r hr ih q.
+    simp ce_prop => {}/ih.
+    case : q.
+    + move => h. inversion 1; subst.
+      inversion H1; subst.
+      sfirstorder use:coqleq_no_hred.
+      have ? : b' = y by eauto using hred_deter. subst.
+      sauto lq:on.
+    + move => h. inversion 1; subst.
+      inversion H0; subst.
+      sfirstorder use:coqleq_no_hred.
+      have ? : b' = y by eauto using hred_deter. subst.
+      sauto lq:on.
+Qed.
diff --git a/theories/termination.v b/theories/termination.v
index 0f8314d..8afe24e 100644
--- a/theories/termination.v
+++ b/theories/termination.v
@@ -107,8 +107,8 @@ Lemma term_metric_pair : forall k a0 b0 a1 b1,
 Proof.
   move => k a0 b0 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
   apply lored_nsteps_pair_inv in hva, hvb.
-  move : hva => [?][?][?][?][?][?][?][?]?.
-  move : hvb => [?][?][?][?][?][?][?][?]?. subst.
+  decompose record hva => {hva}.
+  decompose record hvb => {hvb}. subst.
   simpl in *. exists (k - 1).
   hauto lqb:on solve+:lia.
 Qed.
@@ -119,8 +119,8 @@ Lemma term_metric_bind : forall k p0 a0 b0 p1 a1 b1,
 Proof.
   move => k p0 a0 b0 p1 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
   apply lored_nsteps_bind_inv in hva, hvb.
-  move : hva => [?][?][?][?][?][?][?][?]?.
-  move : hvb => [?][?][?][?][?][?][?][?]?. subst.
+  decompose record hva => {hva}.
+  decompose record hvb => {hvb}. subst.
   simpl in *. exists (k - 1).
   hauto lqb:on solve+:lia.
 Qed.

From 849d19708ee4d1e4eede8236123102e68a22db9c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 14:30:36 -0400
Subject: [PATCH 188/210] Add new definition of algo_dom

---
 theories/algorithmic.v | 362 +++++++++++++++++++-----------------
 theories/common.v      | 412 ++++++++++++++++++++++++++++++++++++++++-
 theories/executable.v  | 212 ---------------------
 theories/termination.v | 261 --------------------------
 4 files changed, 599 insertions(+), 648 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index ab8fe97..dff32da 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -673,8 +673,12 @@ Proof.
     hauto lq:on use:HReds.ToEq, E_Symmetric, E_Transitive.
 Qed.
 
-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.
+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_sym k (a b : PTm) :
+  term_metric k a b -> term_metric k b a.
+Proof. hauto lq:on unfold:term_metric solve+:lia. Qed.
 
 Lemma ne_hne (a : PTm) : ne a -> ishne a.
 Proof. elim : a => //=; sfirstorder b:on. Qed.
@@ -698,37 +702,58 @@ Proof.
   - hauto lq:on use:ne_hne.
 Qed.
 
-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.
+Lemma term_metric_case k (a b : PTm) :
+  term_metric k a b ->
+  (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ term_metric k' a' b /\ k' < k.
 Proof.
-  move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6.
+  move=>[i][j][va][vb][h0] [h1][h2][h3]h4.
   case : a h0 => //=; try firstorder.
   - inversion h0 as [|A B C D E F]; subst.
     hauto qb:on use:ne_hne.
     inversion E; subst => /=.
-    + hauto lq:on use:HRed.AppAbs unfold:algo_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+    + hauto lq:on use:HRed.AppAbs unfold:term_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:term_metric solve+:lia.
     + sfirstorder qb:on use:ne_hne.
   - inversion h0 as [|A B C D E F]; subst.
     hauto qb:on use:ne_hne.
     inversion E; subst => /=.
-    + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
+    + hauto lq:on use:HRed.ProjPair unfold:term_metric solve+:lia.
+    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:term_metric solve+:lia.
   - inversion h0 as [|A B C D E F]; subst.
     hauto qb:on use:ne_hne.
     inversion E; subst => /=.
-    + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + hauto lq:on use:HRed.IndZero unfold:term_metric solve+:lia.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:term_metric solve+:lia.
     + sfirstorder use:ne_hne.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
+    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:term_metric solve+:lia.
     + sfirstorder use:ne_hne.
     + sfirstorder use:ne_hne.
 Qed.
 
-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 A_Conf' a b :
+  ishf a \/ ishne a ->
+  ishf b \/ ishne b ->
+  tm_conf a b ->
+  algo_dom_r a b.
+Proof.
+  move => ha hb.
+  have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
+  have {}hb : HRed.nf b by sfirstorder use:hf_no_hred, hne_no_hred.
+  move => ?.
+  apply A_NfNf.
+  by apply A_Conf.
+Qed.
+
+Lemma hne_nf_ne (a : PTm ) :
+  ishne a -> nf a -> ne a.
+Proof. case : a => //=. Qed.
+
+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 hred_hne (a b : PTm) :
   HRed.R a b ->
@@ -1028,28 +1053,6 @@ Proof.
     exists (S i), a1. hauto lq:on ctrs:nsteps solve+:lia.
 Qed.
 
-Lemma algo_metric_proj k p0 p1 (a0 a1 : PTm) :
-  algo_metric k (PProj p0 a0) (PProj p1 a1) ->
-  ishne a0 ->
-  ishne a1 ->
-  p0 = p1 /\ exists j, j < k /\ algo_metric j a0 a1.
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 hne0 hne1.
-  move : lored_nsteps_proj_inv h0 (hne0);repeat move/[apply].
-  move => [i0][a2][hi][?]ha02. subst.
-  move : lored_nsteps_proj_inv h1 (hne1);repeat move/[apply].
-  move => [i1][a3][hj][?]ha13. subst.
-  simpl in *.
-  move /EJoin.hne_proj_inj : h4 => [h40 h41]. subst.
-  split => //.
-  exists (k - 1). split. simpl in *; lia.
-  rewrite/algo_metric.
-  do 4 eexists. repeat split; eauto. sfirstorder use:ne_nf.
-  sfirstorder use:ne_nf.
-  lia.
-Qed.
-
-
 Lemma hreds_nf_refl a b  :
   HRed.nf a ->
   rtc HRed.R a b ->
@@ -1103,174 +1106,185 @@ Lemma lored_nsteps_pair_inv k (a0 b0 C : PTm ) :
     exists i, (S j), a1, b1. sauto lq:on solve+:lia.
 Qed.
 
-Lemma algo_metric_join k (a b : PTm ) :
-  algo_metric k a b ->
-  DJoin.R a b.
-  rewrite /algo_metric.
-  move => [i][j][va][vb][h0][h1][h2][h3][[v [h40 h41]]]h5.
-  have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
-  have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
-  apply REReds.FromEReds in h40,h41.
-  apply LoReds.ToRReds in h0,h1.
-  apply REReds.FromRReds in h0,h1.
-  rewrite /DJoin.R. exists v. sfirstorder use:@relations.rtc_transitive.
-Qed.
-
-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) ->
-  exists j, j < k /\ algo_metric j a0 a1 /\ algo_metric j b0 b1.
-  move => sn0 sn1 /[dup] /algo_metric_join hj.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
-  move : lored_nsteps_pair_inv h0;repeat move/[apply].
-  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
-  move : lored_nsteps_pair_inv h1;repeat move/[apply].
-  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
+Lemma term_metric_abs : forall k a b,
+    term_metric k (PAbs a) (PAbs b) ->
+    exists k', k' < k /\ term_metric k' a b.
+Proof.
+  move => k a b [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_abs_inv in hva, hvb.
+  move : hva => [a'][hva]?. subst.
+  move : hvb => [b'][hvb]?. subst.
   simpl in *. exists (k - 1).
-  move /andP : h2 => [h20 h21].
-  move /andP : h3 => [h30 h31].
-  suff : EJoin.R a2 a3 /\ EJoin.R b2 b3 by hauto lq:on solve+:lia.
-  hauto l:on use:DJoin.ejoin_pair_inj.
+  hauto lq:on unfold:term_metric solve+:lia.
 Qed.
 
-Lemma hne_nf_ne (a : PTm ) :
-  ishne a -> nf a -> ne a.
-Proof. case : a => //=. Qed.
-
-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).
+Lemma term_metric_pair : forall k a0 b0 a1 b1,
+    term_metric k (PPair a0 b0) (PPair a1 b1) ->
+    exists k', k' < k /\ term_metric k' a0 a1 /\ term_metric k' b0 b1.
 Proof.
-  induction 1; hauto lq:on ctrs:nsteps use:LoRed.renaming.
+  move => k a0 b0 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_pair_inv in hva, hvb.
+  decompose record hva => {hva}.
+  decompose record hvb => {hvb}. subst.
+  simpl in *. exists (k - 1).
+  hauto lqb:on solve+:lia.
 Qed.
 
-Lemma algo_metric_neu_abs k (a0 : PTm) u :
-  algo_metric k u (PAbs a0) ->
+Lemma term_metric_bind : forall k p0 a0 b0 p1 a1 b1,
+    term_metric k (PBind p0 a0 b0) (PBind p1 a1 b1) ->
+    exists k', k' < k /\ term_metric k' a0 a1 /\ term_metric k' b0 b1.
+Proof.
+  move => k p0 a0 b0 p1 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_bind_inv in hva, hvb.
+  decompose record hva => {hva}.
+  decompose record hvb => {hvb}. subst.
+  simpl in *. exists (k - 1).
+  hauto lqb:on solve+:lia.
+Qed.
+
+Lemma term_metric_suc : forall k a b,
+    term_metric k (PSuc a) (PSuc b) ->
+    exists k', k' < k /\ term_metric k' a b.
+Proof.
+  move => k a b [i][j][va][vb][hva][hvb][nfa][nfb]h.
+  apply lored_nsteps_suc_inv in hva, hvb.
+  move : hva => [a'][hva]?. subst.
+  move : hvb => [b'][hvb]?. subst.
+  simpl in *. exists (k - 1).
+  hauto lq:on unfold:term_metric solve+:lia.
+Qed.
+
+Lemma term_metric_abs_neu k (a0 : PTm) u :
+  term_metric k (PAbs a0) u ->
   ishne u ->
-  exists j, j < k /\ algo_metric j (PApp (ren_PTm shift u) (VarPTm var_zero)) a0.
+  exists j, j < k /\ term_metric j a0 (PApp (ren_PTm shift u) (VarPTm var_zero)).
 Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 neu.
-  have neva : ne va by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
-  move /lored_nsteps_abs_inv : h1 => [a1][h01]?. subst.
-  exists (k - 1). simpl in *. split. lia.
-  have {}h0 : nsteps LoRed.R i (ren_PTm shift u) (ren_PTm shift va)
-    by eauto using lored_nsteps_renaming.
-  have {}h0 : nsteps LoRed.R i (PApp (ren_PTm shift u) (VarPTm var_zero)) (PApp (ren_PTm shift va) (VarPTm var_zero)).
-  apply lored_nsteps_app_cong => //=.
-  scongruence use:ishne_ren.
-  do 4 eexists. repeat split; eauto.
-  hauto b:on use:ne_nf_ren.
-  have h : EJoin.R (PAbs (PApp (ren_PTm shift va) (VarPTm var_zero))) (PAbs a1) by hauto q:on ctrs:rtc,ERed.R.
-  apply DJoin.ejoin_abs_inj; eauto.
-  hauto b:on drew:off use:ne_nf_ren.
-  simpl in *.  rewrite size_PTm_ren. lia.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 neu.
+  have neva : ne vb by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
+  move /lored_nsteps_abs_inv : h0 => [a1][h01]?. subst.
+  exists (k - 1).
+  simpl in *. split. lia.
+  exists i,j,a1,(PApp (ren_PTm shift vb) (VarPTm var_zero)).
+  repeat split => //=.
+  apply lored_nsteps_app_cong.
+  by apply lored_nsteps_renaming.
+  by rewrite ishne_ren.
+  rewrite Bool.andb_true_r.
+  sfirstorder use:ne_nf_ren.
+  rewrite size_PTm_ren. lia.
 Qed.
 
-Lemma algo_metric_neu_pair k (a0 b0 : PTm) u :
-  algo_metric k u (PPair a0 b0) ->
+Lemma term_metric_pair_neu k (a0 b0 : PTm) u :
+  term_metric k (PPair a0 b0) u ->
   ishne u ->
-  exists j, j < k /\ algo_metric j (PProj PL u) a0 /\ algo_metric j (PProj PR u) b0.
+  exists j, j < k /\ term_metric j (PProj PL u) a0 /\ term_metric j (PProj PR u) b0.
 Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 neu.
-  move /lored_nsteps_pair_inv : h1.
-  move => [i0][j0][a1][b1][hi][hj][?][ha01]hb01. subst.
-  simpl in *.
-  have ? : ishne va by
-    hauto lq:on use:loreds_hne_preservation, @relations.rtc_nsteps.
-  have ? : ne va by sfirstorder use:hne_nf_ne.
-  exists (k - 1). split. lia.
-  move :lored_nsteps_proj_cong (neu) h0; repeat move/[apply].
-  move => h. have hL := h PL. have {h} hR := h PR.
-  suff [? ?] : EJoin.R (PProj PL va) a1  /\ EJoin.R (PProj PR va) b1.
-  rewrite /algo_metric.
-  split; do 4 eexists; repeat split;eauto; sfirstorder b:on solve+:lia.
-  eapply DJoin.ejoin_pair_inj; hauto qb:on ctrs:rtc, ERed.R.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 neu.
+  have neva : ne vb by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
+  move /lored_nsteps_pair_inv : h0 => [i0][j0][a1][b1][?][?][?][?]?. subst.
+  exists (k-1). sauto qb:on use:lored_nsteps_proj_cong unfold:term_metric solve+:lia.
 Qed.
 
-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) ->
+Lemma term_metric_app k (a0 b0 a1 b1 : PTm) :
+  term_metric k (PApp a0 b0) (PApp a1 b1) ->
   ishne a0 ->
   ishne a1 ->
-  exists j, j < k /\ algo_metric j P0 P1 /\ algo_metric j a0 a1 /\
-         algo_metric j b0 b1 /\ algo_metric j c0 c1.
+  exists j, j < k /\ term_metric j a0 a1 /\ term_metric j b0 b1.
 Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5 hne0 hne1.
-  move /lored_nsteps_ind_inv /(_ hne0) : h0.
-  move =>[iP][ia][ib][ic][P2][a2][b2][c2][?][?][?][?][?][?][?][?]?. subst.
-  move /lored_nsteps_ind_inv /(_ hne1) : h1.
-  move =>[iP0][ia0][ib0][ic0][P3][a3][b3][c3][?][?][?][?][?][?][?][?]?. subst.
-  move /EJoin.ind_inj : h4.
-  move => [?][?][?]?.
-  exists (k -1). simpl in *.
-  hauto lq:on rew:off use:ne_nf b:on solve+:lia.
-Qed.
-
-Lemma algo_metric_app k (a0 b0 a1 b1 : PTm) :
-  algo_metric k (PApp a0 b0) (PApp a1 b1) ->
-  ishne a0 ->
-  ishne a1 ->
-  exists j, j < k /\ algo_metric j a0 a1 /\ algo_metric j b0 b1.
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4.
   move => hne0 hne1.
   move : lored_nsteps_app_inv h0 (hne0);repeat move/[apply].
   move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
   move : lored_nsteps_app_inv h1 (hne1);repeat move/[apply].
   move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
-  simpl in *. exists (k - 1).
-  move /andP : h2 => [h20 h21].
-  move /andP : h3 => [h30 h31].
-  move /EJoin.hne_app_inj : h4 => [h40 h41].
-  split. lia.
-  split.
-  + rewrite /algo_metric.
-    exists i0,j0,a2,a3.
-    repeat split => //=.
-    sfirstorder b:on use:ne_nf.
-    sfirstorder b:on use:ne_nf.
-    lia.
-  + rewrite /algo_metric.
-    exists i1,j1,b2,b3.
-    repeat split => //=; sfirstorder b:on use:ne_nf.
+  simpl in *. exists (k - 1). hauto lqb:on use:lored_nsteps_app_cong, ne_nf unfold:term_metric solve+:lia.
 Qed.
 
-Lemma algo_metric_suc k (a0 a1 : PTm) :
-  algo_metric k (PSuc a0) (PSuc a1) ->
-  exists j, j < k /\ algo_metric j a0 a1.
+Lemma term_metric_proj k p0 p1 (a0 a1 : PTm) :
+  term_metric k (PProj p0 a0) (PProj p1 a1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ term_metric j a0 a1.
 Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
-  exists (k - 1).
-  move /lored_nsteps_suc_inv : h0.
-  move => [a0'][ha0]?. subst.
-  move /lored_nsteps_suc_inv : h1.
-  move => [b0'][hb0]?. subst. simpl in *.
-  split; first by lia.
-  rewrite /algo_metric.
-  hauto lq:on rew:off use:EJoin.suc_inj solve+:lia.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 hne0 hne1.
+  move : lored_nsteps_proj_inv h0 (hne0);repeat move/[apply].
+  move => [i0][a2][hi][?]ha02. subst.
+  move : lored_nsteps_proj_inv h1 (hne1);repeat move/[apply].
+  move => [i1][a3][hj][?]ha13. subst.
+  exists (k- 1).  hauto q:on use:ne_nf solve+:lia.
 Qed.
 
-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.
+Lemma term_metric_ind k P0 (a0 : PTm ) b0 c0 P1 a1 b1 c1 :
+  term_metric k (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
+  ishne a0 ->
+  ishne a1 ->
+  exists j, j < k /\ term_metric j P0 P1 /\ term_metric j a0 a1 /\
+         term_metric j b0 b1 /\ term_metric j c0 c1.
 Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
-  move : lored_nsteps_bind_inv h0 => /[apply].
-  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
-  move : lored_nsteps_bind_inv h1 => /[apply].
-  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
-  move /EJoin.bind_inj : h4 => [?][hEa]hEb. subst.
-  split => //.
-  exists (k - 1). split. simpl in *; lia.
-  simpl in *.
-  split.
-  - exists i0,j0,a2,a3.
-    repeat split => //=; sfirstorder b:on solve+:lia.
-  - exists i1,j1,b2,b3. sfirstorder b:on solve+:lia.
+  move => [i][j][va][vb][h0][h1][h2][h3]h4 hne0 hne1.
+  move /lored_nsteps_ind_inv /(_ hne0) : h0.
+  move =>[iP][ia][ib][ic][P2][a2][b2][c2][?][?][?][?][?][?][?][?]?. subst.
+  move /lored_nsteps_ind_inv /(_ hne1) : h1.
+  move =>[iP0][ia0][ib0][ic0][P3][a3][b3][c3][?][?][?][?][?][?][?][?]?. subst.
+  exists (k -1). simpl in *.
+  hauto lq:on rew:off use:ne_nf b:on solve+:lia.
 Qed.
 
 
+Lemma algo_dom_r_algo_dom : forall a b, HRed.nf a -> HRed.nf b -> algo_dom_r a b -> algo_dom a b.
+Proof. hauto l:on use:algo_dom_r_inv, hreds_nf_refl. Qed.
+
+Lemma term_metric_algo_dom : forall k a b, term_metric k a b -> algo_dom_r a b.
+Proof.
+  move => [:hneL].
+  elim /Wf_nat.lt_wf_ind.
+  move => n ih a b h.
+  case /term_metric_case : (h); cycle 1.
+  move => [k'][a'][h0][h1]h2.
+  by apply : A_HRedL; eauto.
+  case /term_metric_sym /term_metric_case : (h); cycle 1.
+  move => [k'][b'][hb][/term_metric_sym h0]h1.
+  move => ha. have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
+  by apply : A_HRedR; eauto.
+  move => /[swap].
+  case => hfa; case => hfb.
+  - move : hfa hfb h.
+    case : a; case : b => //=; eauto 5 using A_Conf' with adom.
+    + hauto lq:on use:term_metric_abs db:adom.
+    + hauto lq:on use:term_metric_pair db:adom.
+    + hauto lq:on use:term_metric_bind db:adom.
+    + hauto lq:on use:term_metric_suc db:adom.
+  - abstract : hneL n ih a b h hfa hfb.
+    case : a hfa h => //=.
+    + hauto lq:on use:term_metric_abs_neu db:adom.
+    + scrush use:term_metric_pair_neu db:adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
+  - hauto q:on use:algo_dom_sym, term_metric_sym.
+  - move {hneL}.
+    case : b hfa hfb h => //=; case a => //=; eauto 5 using A_Conf' with adom.
+    + move => a0 b0 a1 b1 nfa0 nfa1.
+      move /term_metric_app /(_ nfa0  nfa1) => [j][hj][ha]hb.
+      apply A_NfNf. apply A_AppCong => //; eauto.
+      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
+      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
+      eauto using algo_dom_r_algo_dom.
+    + move => p0 A0 p1 A1 neA0 neA1.
+      have {}nfa0 : HRed.nf A0 by sfirstorder use:hne_no_hred.
+      have {}nfb0 : HRed.nf A1 by sfirstorder use:hne_no_hred.
+      hauto lq:on use:term_metric_proj, algo_dom_r_algo_dom db:adom.
+    + move => P0 a0 b0 c0 P1 a1 b1 c1 nea0 nea1.
+      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
+      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
+      hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
+Qed.
+
 Lemma coqeq_complete' k (a b : PTm ) :
+  (forall a b, algo_dom a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b)
+
   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).
diff --git a/theories/common.v b/theories/common.v
index 3ea15d6..c14bbb3 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.unscoped Autosubst2.syntax Autosubst2.core ssreflect.
+Require Import Autosubst2.unscoped Autosubst2.syntax Autosubst2.core ssreflect ssrbool.
 From Equations Require Import Equations.
 Derive NoConfusion for nat PTag BTag PTm.
 Derive EqDec for BTag PTag PTm.
@@ -6,6 +6,7 @@ From Ltac2 Require Ltac2.
 Import Ltac2.Notations.
 Import Ltac2.Control.
 From Hammer Require Import Tactics.
+From stdpp Require Import relations (rtc(..)).
 
 Inductive lookup : nat -> list PTm -> PTm -> Prop :=
 | here A Γ : lookup 0 (cons A Γ) (ren_PTm shift A)
@@ -119,6 +120,28 @@ Definition isabs (a : PTm) :=
   | _ => false
   end.
 
+Definition tm_nonconf (a b : PTm) : bool :=
+  match a, b with
+  | PAbs _, _ => (~~ ishf b) || isabs b
+  | _, PAbs _ => ~~ ishf a
+  | VarPTm _, VarPTm _ => true
+  | PPair _ _, _ => (~~ ishf b) || ispair b
+  | _, PPair _ _ => ~~ ishf a
+  | PZero, PZero => true
+  | PSuc _, PSuc _ => true
+  | 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
+  | _,_=> false
+  end.
+
+Definition tm_conf (a b : PTm) := ~~ tm_nonconf a b.
+
+
+
 Definition ishf_ren (a : PTm)  (ξ : nat -> nat) :
   ishf (ren_PTm ξ a) = ishf a.
 Proof. case : a => //=. Qed.
@@ -175,3 +198,390 @@ Module HRed.
 
     Definition nf a := forall b, ~ R a b.
 End HRed.
+
+Inductive algo_dom : PTm -> PTm -> Prop :=
+| A_AbsAbs a b :
+  algo_dom_r a b ->
+  (* --------------------- *)
+  algo_dom (PAbs a) (PAbs b)
+
+| A_AbsNeu a u :
+  ishne u ->
+  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_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a ->
+  (* --------------------- *)
+  algo_dom u (PAbs a)
+
+| A_PairPair a0 a1 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_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_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_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_NeuNeu a b :
+  algo_dom_neu 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_neu : PTm -> PTm -> Prop :=
+| A_VarCong i j :
+  (* -------------------------- *)
+  algo_dom_neu (VarPTm i) (VarPTm j)
+
+| A_AppCong u0 u1 a0 a1 :
+  ishne u0 ->
+  ishne u1 ->
+  algo_dom_neu u0 u1 ->
+  algo_dom_r a0 a1 ->
+  (* ------------------------- *)
+  algo_dom_neu (PApp u0 a0) (PApp u1 a1)
+
+| A_ProjCong p0 p1 u0 u1 :
+  ishne u0 ->
+  ishne u1 ->
+  algo_dom_neu u0 u1 ->
+  (* ---------------------  *)
+  algo_dom_neu (PProj p0 u0) (PProj p1 u1)
+
+| A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 :
+  ishne u0 ->
+  ishne u1 ->
+  algo_dom_r P0 P1 ->
+  algo_dom_neu u0 u1 ->
+  algo_dom_r b0 b1 ->
+  algo_dom_r c0 c1 ->
+  algo_dom_neu (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_r a' b ->
+  (* ----------------------- *)
+  algo_dom_r a b
+
+| A_HRedR a b b'  :
+  HRed.nf a ->
+  HRed.R b b' ->
+  algo_dom_r a b' ->
+  (* ----------------------- *)
+  algo_dom_r a b.
+
+Scheme algo_ind := Induction for algo_dom Sort Prop
+  with algo_neu_ind := Induction for algo_dom_neu Sort Prop
+  with algor_ind := Induction for algo_dom_r Sort Prop.
+
+Combined Scheme algo_dom_mutual from algo_ind, algo_neu_ind, algor_ind.
+#[export]Hint Constructors algo_dom algo_dom_neu algo_dom_r : adom.
+
+Definition stm_nonconf a b :=
+  match a, b with
+  | PUniv _, PUniv _ => true
+  | PBind PPi _ _, PBind PPi _ _ => true
+  | PBind PSig _ _, PBind PSig _ _ => true
+  | PNat, PNat => true
+  | VarPTm _, VarPTm _ => true
+  | PApp _ _, PApp _ _ => (~~ ishf a) && (~~ ishf b)
+  | PProj _ _, PProj _ _ => (~~ ishf a) && (~~ ishf b)
+  | PInd _ _ _ _, PInd _ _ _ _ => (~~ ishf a) && (~~ ishf b)
+  | _, _ => false
+  end.
+
+Lemma stm_tm_nonconf a b :
+  stm_nonconf a b -> tm_nonconf a b.
+Proof. apply /implyP.
+       destruct a ,b =>//=; hauto lq:on inv:PTag, BTag.
+Qed.
+
+Definition stm_conf a b := ~~ stm_nonconf a b.
+
+Lemma tm_stm_conf a b :
+  tm_conf a b -> stm_conf a b.
+Proof.
+  rewrite /tm_conf /stm_conf.
+  apply /contra /stm_tm_nonconf.
+Qed.
+
+Inductive salgo_dom : PTm -> PTm -> Prop :=
+| S_UnivCong i j :
+  (* -------------------------- *)
+  salgo_dom (PUniv i) (PUniv j)
+
+| S_PiCong  A0 A1 B0 B1 :
+  salgo_dom_r A1 A0 ->
+  salgo_dom_r B0 B1 ->
+  (* ---------------------------- *)
+  salgo_dom (PBind PPi A0 B0) (PBind PPi A1 B1)
+
+| S_SigCong  A0 A1 B0 B1 :
+  salgo_dom_r A0 A1 ->
+  salgo_dom_r B0 B1 ->
+  (* ---------------------------- *)
+  salgo_dom (PBind PSig A0 B0) (PBind PSig A1 B1)
+
+| S_NatCong :
+  salgo_dom PNat PNat
+
+| S_NeuNeu a b :
+  algo_dom_neu a b ->
+  salgo_dom a b
+
+| S_Conf a b :
+  HRed.nf a ->
+  HRed.nf b ->
+  stm_conf a b ->
+  salgo_dom a b
+
+with salgo_dom_r : PTm -> PTm -> Prop :=
+| S_NfNf a b :
+  salgo_dom a b ->
+  salgo_dom_r a b
+
+| S_HRedL a a' b  :
+  HRed.R a a' ->
+  salgo_dom_r a' b ->
+  (* ----------------------- *)
+  salgo_dom_r a b
+
+| S_HRedR a b b'  :
+  HRed.nf a ->
+  HRed.R b b' ->
+  salgo_dom_r a b' ->
+  (* ----------------------- *)
+  salgo_dom_r a b.
+
+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.
+
+Ltac2 destruct_salgo () :=
+  lazy_match! goal with
+  | [h : is_true (ishne ?a) |- is_true (stm_conf ?a _) ] =>
+      if Constr.is_var a then destruct $a; ltac1:(done) else ()
+  | [|- is_true (stm_conf _ _)] =>
+      unfold stm_conf; ltac1:(done)
+  end.
+
+Ltac destruct_salgo := ltac2:(destruct_salgo ()).
+
+Lemma algo_dom_r_inv a b :
+  algo_dom_r a b -> exists a0 b0, algo_dom a0 b0 /\ rtc HRed.R a a0 /\ rtc HRed.R b b0.
+Proof.
+  induction 1; hauto lq:on ctrs:rtc.
+Qed.
+
+Lemma A_HRedsL a a' b :
+  rtc HRed.R a a' ->
+  algo_dom_r a' b ->
+  algo_dom_r a b.
+  induction 1; sfirstorder use:A_HRedL.
+Qed.
+
+Lemma A_HRedsR a b b' :
+  HRed.nf a ->
+  rtc HRed.R b b' ->
+  algo_dom a b' ->
+  algo_dom_r a b.
+Proof. induction 2; sauto. Qed.
+
+Lemma tm_conf_sym a b : tm_conf a b = tm_conf b a.
+Proof. case : a; case : b => //=. Qed.
+
+Lemma algo_dom_neu_hne (a b : PTm) :
+  algo_dom_neu a b ->
+  ishne a /\ ishne b.
+Proof. inversion 1; subst => //=. 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 use:algo_dom_neu_hne lq:on unfold:HRed.nf.
+Qed.
+
+
+Lemma A_HRedR' a b b' :
+  HRed.R b b' ->
+  algo_dom_r a b' ->
+  algo_dom_r a b.
+Proof.
+  move => hb /algo_dom_r_inv.
+  move => [a0 [b0 [h0 [h1 h2]]]].
+  have {h2} {}hb : rtc HRed.R b b0 by hauto lq:on ctrs:rtc.
+  have ? : HRed.nf a0 by sfirstorder use:algo_dom_no_hred.
+  hauto lq:on use:A_HRedsL, A_HRedsR.
+Qed.
+
+Lemma algo_dom_sym :
+  (forall a b (h : algo_dom a b), algo_dom b a) /\
+    (forall a b, algo_dom_neu a b -> algo_dom_neu b a) /\
+    (forall a b (h : algo_dom_r a b), algo_dom_r b a).
+Proof.
+  apply algo_dom_mutual; try qauto use:tm_conf_sym,A_HRedR' db:adom.
+Qed.
+
+Lemma salgo_dom_r_inv a b :
+  salgo_dom_r a b -> exists a0 b0, salgo_dom a0 b0 /\ rtc HRed.R a a0 /\ rtc HRed.R b b0.
+Proof.
+  induction 1; hauto lq:on ctrs:rtc.
+Qed.
+
+Lemma S_HRedsL a a' b :
+  rtc HRed.R a a' ->
+  salgo_dom_r a' b ->
+  salgo_dom_r a b.
+  induction 1; sfirstorder use:S_HRedL.
+Qed.
+
+Lemma S_HRedsR a b b' :
+  HRed.nf a ->
+  rtc HRed.R b b' ->
+  salgo_dom a b' ->
+  salgo_dom_r a b.
+Proof. induction 2; sauto. Qed.
+
+Lemma stm_conf_sym a b : stm_conf a b = stm_conf b a.
+Proof. case : a; case : b => //=; hauto lq:on inv:PTag, BTag. Qed.
+
+Lemma salgo_dom_no_hred (a b : PTm) :
+  salgo_dom a b ->
+  HRed.nf a /\ HRed.nf b.
+Proof.
+  induction 1 =>//=; try hauto inv:HRed.R use:hne_no_hred, hf_no_hred, algo_dom_neu_hne lq:on unfold:HRed.nf.
+Qed.
+
+Lemma S_HRedR' a b b' :
+  HRed.R b b' ->
+  salgo_dom_r a b' ->
+  salgo_dom_r a b.
+Proof.
+  move => hb /salgo_dom_r_inv.
+  move => [a0 [b0 [h0 [h1 h2]]]].
+  have {h2} {}hb : rtc HRed.R b b0 by hauto lq:on ctrs:rtc.
+  have ? : HRed.nf a0 by sfirstorder use:salgo_dom_no_hred.
+  hauto lq:on use:S_HRedsL, S_HRedsR.
+Qed.
+
+Lemma algo_dom_salgo_dom :
+  (forall a b, algo_dom a b -> salgo_dom a b /\ salgo_dom b a) /\
+    (forall a b, algo_dom_neu a b -> True) /\
+    (forall a b, algo_dom_r a b -> salgo_dom_r a b /\ salgo_dom_r b a).
+Proof.
+  apply algo_dom_mutual => //=;
+                             try hauto lq:on ctrs:salgo_dom, algo_dom_neu, salgo_dom_r use:S_Conf, hne_no_hred, algo_dom_sym, tm_stm_conf, S_HRedR' inv:HRed.R unfold:HRed.nf  solve+:destruct_salgo.
+  - case;case; hauto lq:on ctrs:salgo_dom, algo_dom_neu, salgo_dom_r use:S_Conf, hne_no_hred, algo_dom_sym inv:HRed.R unfold:HRed.nf  solve+:destruct_salgo.
+  - move => a b ha hb /[dup] /tm_stm_conf h.
+    rewrite tm_conf_sym => /tm_stm_conf h0.
+    hauto l:on use:S_Conf inv:HRed.R unfold:HRed.nf.
+Qed.
+
+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.
diff --git a/theories/executable.v b/theories/executable.v
index 44557f9..fd03f2d 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -11,26 +11,6 @@ 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 _, _ => (~~ ishf b) || isabs b
-  | _, PAbs _ => ~~ ishf a
-  | VarPTm _, VarPTm _ => true
-  | PPair _ _, _ => (~~ ishf b) || ispair b
-  | _, PPair _ _ => ~~ ishf a
-  | PZero, PZero => true
-  | PSuc _, PSuc _ => true
-  | 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
-  | _,_=> 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)
@@ -102,121 +82,6 @@ Equations tm_to_eq_view (a b : PTm) : eq_view a b :=
   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_r a b ->
-  (* --------------------- *)
-  algo_dom (PAbs a) (PAbs b)
-
-| A_AbsNeu a u :
-  ishne u ->
-  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_r (PApp (ren_PTm shift u) (VarPTm var_zero)) a ->
-  (* --------------------- *)
-  algo_dom u (PAbs a)
-
-| A_PairPair a0 a1 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_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_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_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)
-
-| A_ProjCong p0 p1 u0 u1 :
-  ishne u0 ->
-  ishne u1 ->
-  algo_dom u0 u1 ->
-  (* ---------------------  *)
-  algo_dom (PProj p0 u0) (PProj p1 u1)
-
-| A_AppCong u0 u1 a0 a1 :
-  ishne u0 ->
-  ishne u1 ->
-  algo_dom u0 u1 ->
-  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)
-
-| 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 :
-  algo_dom a b ->
-  algo_dom_r a b
-
-| A_HRedL a a' b  :
-  HRed.R a a' ->
-  algo_dom_r a' b ->
-  (* ----------------------- *)
-  algo_dom_r a b
-
-| A_HRedR a b b'  :
-  HRed.nf a ->
-  HRed.R b b' ->
-  algo_dom_r a b' ->
-  (* ----------------------- *)
-  algo_dom_r a b.
-
-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.
-
 
 (* Inductive salgo_dom : PTm -> PTm -> Prop := *)
 (* | S_UnivCong i j : *)
@@ -273,83 +138,6 @@ Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
 (*   with algor_ind := Induction for salgo_dom_r Sort Prop. *)
 
 
-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.
-
-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) :=
-    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.
-
 Ltac2 destruct_algo () :=
   lazy_match! goal with
   | [h : algo_dom ?a ?b |- _ ] =>
diff --git a/theories/termination.v b/theories/termination.v
index 8afe24e..41400ef 100644
--- a/theories/termination.v
+++ b/theories/termination.v
@@ -4,267 +4,6 @@ Require Import ssreflect ssrbool.
 From stdpp Require Import relations (nsteps (..), rtc(..)).
 Import Psatz.
 
-Lemma tm_conf_sym a b : tm_conf a b = tm_conf b a.
-Proof. case : a; case : b => //=. Qed.
-
-#[export]Hint Constructors algo_dom algo_dom_r : adom.
-
-Lemma algo_dom_r_inv a b :
-  algo_dom_r a b -> exists a0 b0, algo_dom a0 b0 /\ rtc HRed.R a a0 /\ rtc HRed.R b b0.
-Proof.
-  induction 1; hauto lq:on ctrs:rtc.
-Qed.
-
-Lemma A_HRedsL a a' b :
-  rtc HRed.R a a' ->
-  algo_dom_r a' b ->
-  algo_dom_r a b.
-  induction 1; sfirstorder use:A_HRedL.
-Qed.
-
-Lemma A_HRedsR a b b' :
-  HRed.nf a ->
-  rtc HRed.R b b' ->
-  algo_dom a b' ->
-  algo_dom_r a b.
-Proof. induction 2; sauto. Qed.
-
-Lemma algo_dom_sym :
-  (forall a b (h : algo_dom a b), algo_dom b a) /\
-    (forall a b (h : algo_dom_r a b), algo_dom_r b a).
-Proof.
-  apply algo_dom_mutual; try qauto use:tm_conf_sym db:adom.
-  move => a a' b hr h ih.
-  move /algo_dom_r_inv : ih => [a0][b0][ih0][ih1]ih2.
-  have nfa0 : HRed.nf a0 by sfirstorder use:algo_dom_no_hred.
-  sauto use:A_HRedsL, A_HRedsR.
-Qed.
-
-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_sym k (a b : PTm) :
-  term_metric k a b -> term_metric k b a.
-Proof. hauto lq:on unfold:term_metric solve+:lia. Qed.
-
-Lemma term_metric_case k (a b : PTm) :
-  term_metric k a b ->
-  (ishf a \/ ishne a) \/ exists k' a', HRed.R a a' /\ term_metric k' a' b /\ k' < k.
-Proof.
-  move=>[i][j][va][vb][h0] [h1][h2][h3]h4.
-  case : a h0 => //=; try firstorder.
-  - inversion h0 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.AppAbs unfold:term_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:term_metric solve+:lia.
-    + sfirstorder qb:on use:ne_hne.
-  - inversion h0 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.ProjPair unfold:term_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:term_metric solve+:lia.
-  - inversion h0 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.IndZero unfold:term_metric solve+:lia.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:term_metric solve+:lia.
-    + sfirstorder use:ne_hne.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:term_metric solve+:lia.
-    + sfirstorder use:ne_hne.
-    + sfirstorder use:ne_hne.
-Qed.
-
-Lemma A_Conf' a b :
-  ishf a \/ ishne a ->
-  ishf b \/ ishne b ->
-  tm_conf a b ->
-  algo_dom_r a b.
-Proof.
-  move => ha hb.
-  have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
-  have {}hb : HRed.nf b by sfirstorder use:hf_no_hred, hne_no_hred.
-  move => ?.
-  apply A_NfNf.
-  by apply A_Conf.
-Qed.
-
-Lemma term_metric_abs : forall k a b,
-    term_metric k (PAbs a) (PAbs b) ->
-    exists k', k' < k /\ term_metric k' a b.
-Proof.
-  move => k a b [i][j][va][vb][hva][hvb][nfa][nfb]h.
-  apply lored_nsteps_abs_inv in hva, hvb.
-  move : hva => [a'][hva]?. subst.
-  move : hvb => [b'][hvb]?. subst.
-  simpl in *. exists (k - 1).
-  hauto lq:on unfold:term_metric solve+:lia.
-Qed.
-
-Lemma term_metric_pair : forall k a0 b0 a1 b1,
-    term_metric k (PPair a0 b0) (PPair a1 b1) ->
-    exists k', k' < k /\ term_metric k' a0 a1 /\ term_metric k' b0 b1.
-Proof.
-  move => k a0 b0 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
-  apply lored_nsteps_pair_inv in hva, hvb.
-  decompose record hva => {hva}.
-  decompose record hvb => {hvb}. subst.
-  simpl in *. exists (k - 1).
-  hauto lqb:on solve+:lia.
-Qed.
-
-Lemma term_metric_bind : forall k p0 a0 b0 p1 a1 b1,
-    term_metric k (PBind p0 a0 b0) (PBind p1 a1 b1) ->
-    exists k', k' < k /\ term_metric k' a0 a1 /\ term_metric k' b0 b1.
-Proof.
-  move => k p0 a0 b0 p1 a1 b1 [i][j][va][vb][hva][hvb][nfa][nfb]h.
-  apply lored_nsteps_bind_inv in hva, hvb.
-  decompose record hva => {hva}.
-  decompose record hvb => {hvb}. subst.
-  simpl in *. exists (k - 1).
-  hauto lqb:on solve+:lia.
-Qed.
-
-Lemma term_metric_suc : forall k a b,
-    term_metric k (PSuc a) (PSuc b) ->
-    exists k', k' < k /\ term_metric k' a b.
-Proof.
-  move => k a b [i][j][va][vb][hva][hvb][nfa][nfb]h.
-  apply lored_nsteps_suc_inv in hva, hvb.
-  move : hva => [a'][hva]?. subst.
-  move : hvb => [b'][hvb]?. subst.
-  simpl in *. exists (k - 1).
-  hauto lq:on unfold:term_metric solve+:lia.
-Qed.
-
-Lemma term_metric_abs_neu k (a0 : PTm) u :
-  term_metric k (PAbs a0) u ->
-  ishne u ->
-  exists j, j < k /\ term_metric j a0 (PApp (ren_PTm shift u) (VarPTm var_zero)).
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3]h4 neu.
-  have neva : ne vb by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
-  move /lored_nsteps_abs_inv : h0 => [a1][h01]?. subst.
-  exists (k - 1).
-  simpl in *. split. lia.
-  exists i,j,a1,(PApp (ren_PTm shift vb) (VarPTm var_zero)).
-  repeat split => //=.
-  apply lored_nsteps_app_cong.
-  by apply lored_nsteps_renaming.
-  by rewrite ishne_ren.
-  rewrite Bool.andb_true_r.
-  sfirstorder use:ne_nf_ren.
-  rewrite size_PTm_ren. lia.
-Qed.
-
-Lemma term_metric_pair_neu k (a0 b0 : PTm) u :
-  term_metric k (PPair a0 b0) u ->
-  ishne u ->
-  exists j, j < k /\ term_metric j (PProj PL u) a0 /\ term_metric j (PProj PR u) b0.
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3]h4 neu.
-  have neva : ne vb by hauto lq:on use:hne_nf_ne, loreds_hne_preservation, @relations.rtc_nsteps.
-  move /lored_nsteps_pair_inv : h0 => [i0][j0][a1][b1][?][?][?][?]?. subst.
-  exists (k-1). sauto qb:on use:lored_nsteps_proj_cong unfold:term_metric solve+:lia.
-Qed.
-
-Lemma term_metric_app k (a0 b0 a1 b1 : PTm) :
-  term_metric k (PApp a0 b0) (PApp a1 b1) ->
-  ishne a0 ->
-  ishne a1 ->
-  exists j, j < k /\ term_metric j a0 a1 /\ term_metric j b0 b1.
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3]h4.
-  move => hne0 hne1.
-  move : lored_nsteps_app_inv h0 (hne0);repeat move/[apply].
-  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
-  move : lored_nsteps_app_inv h1 (hne1);repeat move/[apply].
-  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
-  simpl in *. exists (k - 1). hauto lqb:on use:lored_nsteps_app_cong, ne_nf unfold:term_metric solve+:lia.
-Qed.
-
-Lemma term_metric_proj k p0 p1 (a0 a1 : PTm) :
-  term_metric k (PProj p0 a0) (PProj p1 a1) ->
-  ishne a0 ->
-  ishne a1 ->
-  exists j, j < k /\ term_metric j a0 a1.
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3]h4 hne0 hne1.
-  move : lored_nsteps_proj_inv h0 (hne0);repeat move/[apply].
-  move => [i0][a2][hi][?]ha02. subst.
-  move : lored_nsteps_proj_inv h1 (hne1);repeat move/[apply].
-  move => [i1][a3][hj][?]ha13. subst.
-  exists (k- 1).  hauto q:on use:ne_nf solve+:lia.
-Qed.
-
-Lemma term_metric_ind k P0 (a0 : PTm ) b0 c0 P1 a1 b1 c1 :
-  term_metric k (PInd P0 a0 b0 c0) (PInd P1 a1 b1 c1) ->
-  ishne a0 ->
-  ishne a1 ->
-  exists j, j < k /\ term_metric j P0 P1 /\ term_metric j a0 a1 /\
-         term_metric j b0 b1 /\ term_metric j c0 c1.
-Proof.
-  move => [i][j][va][vb][h0][h1][h2][h3]h4 hne0 hne1.
-  move /lored_nsteps_ind_inv /(_ hne0) : h0.
-  move =>[iP][ia][ib][ic][P2][a2][b2][c2][?][?][?][?][?][?][?][?]?. subst.
-  move /lored_nsteps_ind_inv /(_ hne1) : h1.
-  move =>[iP0][ia0][ib0][ic0][P3][a3][b3][c3][?][?][?][?][?][?][?][?]?. subst.
-  exists (k -1). simpl in *.
-  hauto lq:on rew:off use:ne_nf b:on solve+:lia.
-Qed.
-
-
-Lemma algo_dom_r_algo_dom : forall a b, HRed.nf a -> HRed.nf b -> algo_dom_r a b -> algo_dom a b.
-Proof. hauto l:on use:algo_dom_r_inv, hreds_nf_refl. Qed.
-
-Lemma term_metric_algo_dom : forall k a b, term_metric k a b -> algo_dom_r a b.
-Proof.
-  move => [:hneL].
-  elim /Wf_nat.lt_wf_ind.
-  move => n ih a b h.
-  case /term_metric_case : (h); cycle 1.
-  move => [k'][a'][h0][h1]h2.
-  by apply : A_HRedL; eauto.
-  case /term_metric_sym /term_metric_case : (h); cycle 1.
-  move => [k'][b'][hb][/term_metric_sym h0]h1.
-  move => ha. have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
-  by apply : A_HRedR; eauto.
-  move => /[swap].
-  case => hfa; case => hfb.
-  - move : hfa hfb h.
-    case : a; case : b => //=; eauto 5 using A_Conf' with adom.
-    + hauto lq:on use:term_metric_abs db:adom.
-    + hauto lq:on use:term_metric_pair db:adom.
-    + hauto lq:on use:term_metric_bind db:adom.
-    + hauto lq:on use:term_metric_suc db:adom.
-  - abstract : hneL n ih a b h hfa hfb.
-    case : a hfa h => //=.
-    + hauto lq:on use:term_metric_abs_neu db:adom.
-    + scrush use:term_metric_pair_neu db:adom.
-    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
-    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
-    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
-    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
-    + case : b hfb => //=; eauto 5 using A_Conf' with adom.
-  - hauto q:on use:algo_dom_sym, term_metric_sym.
-  - move {hneL}.
-    case : b hfa hfb h => //=; case a => //=; eauto 5 using A_Conf' with adom.
-    + move => a0 b0 a1 b1 nfa0 nfa1.
-      move /term_metric_app /(_ nfa0  nfa1) => [j][hj][ha]hb.
-      apply A_NfNf. apply A_AppCong => //; eauto.
-      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
-      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
-      eauto using algo_dom_r_algo_dom.
-    + move => p0 A0 p1 A1 neA0 neA1.
-      have {}nfa0 : HRed.nf A0 by sfirstorder use:hne_no_hred.
-      have {}nfb0 : HRed.nf A1 by sfirstorder use:hne_no_hred.
-      hauto lq:on use:term_metric_proj, algo_dom_r_algo_dom db:adom.
-    + move => P0 a0 b0 c0 P1 a1 b1 c1 nea0 nea1.
-      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
-      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
-      hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
-Qed.
-
 Lemma sn_term_metric (a b : PTm) : SN a -> SN b -> exists k, term_metric k a b.
 Proof.
   move /LoReds.FromSN => [va [ha0 ha1]].

From 4cbd2ac0fd5449b27ecc3c50b3aa11b9dbd7234b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 15:35:43 -0400
Subject: [PATCH 189/210] Save

---
 theories/algorithmic.v | 16 +++++++++++-----
 theories/common.v      | 18 +++++++++++-------
 2 files changed, 22 insertions(+), 12 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index dff32da..17868a0 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -737,8 +737,6 @@ Lemma A_Conf' a b :
   algo_dom_r a b.
 Proof.
   move => ha hb.
-  have {}ha : HRed.nf a by sfirstorder use:hf_no_hred, hne_no_hred.
-  have {}hb : HRed.nf b by sfirstorder use:hf_no_hred, hne_no_hred.
   move => ?.
   apply A_NfNf.
   by apply A_Conf.
@@ -1234,6 +1232,11 @@ Qed.
 Lemma algo_dom_r_algo_dom : forall a b, HRed.nf a -> HRed.nf b -> algo_dom_r a b -> algo_dom a b.
 Proof. hauto l:on use:algo_dom_r_inv, hreds_nf_refl. Qed.
 
+Lemma algo_dom_algo_dom_neu : forall a b, ishne a -> ishne b -> algo_dom a b -> algo_dom_neu a b \/ tm_conf a b.
+Proof.
+  inversion 3; subst => //=; tauto.
+Qed.
+
 Lemma term_metric_algo_dom : forall k a b, term_metric k a b -> algo_dom_r a b.
 Proof.
   move => [:hneL].
@@ -1268,9 +1271,12 @@ Proof.
     case : b hfa hfb h => //=; case a => //=; eauto 5 using A_Conf' with adom.
     + move => a0 b0 a1 b1 nfa0 nfa1.
       move /term_metric_app /(_ nfa0  nfa1) => [j][hj][ha]hb.
-      apply A_NfNf. apply A_AppCong => //; eauto.
-      have {}nfa0 : HRed.nf a0 by sfirstorder use:hne_no_hred.
-      have {}nfb0 : HRed.nf a1 by sfirstorder use:hne_no_hred.
+      apply A_NfNf.
+      (* apply A_NfNf. apply A_NeuNeu. apply A_AppCong => //; eauto. *)
+      have nfa0' : HRed.nf a0 by sfirstorder use:hne_no_hred.
+      have nfb0' : HRed.nf a1 by sfirstorder use:hne_no_hred.
+      have ha0 : algo_dom a0 a1 by eauto using algo_dom_r_algo_dom.
+      apply A_Conf. admit. admit. rewrite /tm_conf. simpl.
       eauto using algo_dom_r_algo_dom.
     + move => p0 A0 p1 A1 neA0 neA1.
       have {}nfa0 : HRed.nf A0 by sfirstorder use:hne_no_hred.
diff --git a/theories/common.v b/theories/common.v
index c14bbb3..b4867d2 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -129,9 +129,9 @@ Definition tm_nonconf (a b : PTm) : bool :=
   | _, PPair _ _ => ~~ ishf a
   | PZero, PZero => true
   | PSuc _, PSuc _ => true
-  | PApp _ _, PApp _ _ => (~~ ishf a) && (~~ ishf b)
-  | PProj _ _, PProj _ _ => (~~ ishf a) && (~~ ishf b)
-  | PInd _ _ _ _, PInd _ _ _ _ => (~~ ishf a) && (~~ ishf b)
+  | PApp _ _, PApp _ _ => true
+  | PProj _ _, PProj _ _ => true
+  | PInd _ _ _ _, PInd _ _ _ _ => true
   | PNat, PNat => true
   | PUniv _, PUniv _ => true
   | PBind _ _ _, PBind _ _ _ => true
@@ -140,8 +140,6 @@ Definition tm_nonconf (a b : PTm) : bool :=
 
 Definition tm_conf (a b : PTm) := ~~ tm_nonconf a b.
 
-
-
 Definition ishf_ren (a : PTm)  (ξ : nat -> nat) :
   ishf (ren_PTm ξ a) = ishf a.
 Proof. case : a => //=. Qed.
@@ -262,8 +260,8 @@ Inductive algo_dom : PTm -> PTm -> Prop :=
   algo_dom a b
 
 | A_Conf a b :
-  HRed.nf a ->
-  HRed.nf b ->
+  ishf a ->
+  ishf b ->
   tm_conf a b ->
   algo_dom a b
 
@@ -296,6 +294,12 @@ with algo_dom_neu : PTm -> PTm -> Prop :=
   algo_dom_r c0 c1 ->
   algo_dom_neu (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
 
+| A_NeuConf a b :
+  ishne a ->
+  ishne b ->
+  tm_conf a b ->
+  algo_dom_neu a b
+
 with algo_dom_r : PTm  -> PTm -> Prop :=
 | A_NfNf a b :
   algo_dom a b ->

From e278c6eaefcfacd9a178b0c1d7fb4d1b681eb3a0 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 16:22:37 -0400
Subject: [PATCH 190/210] Define salgo_dom

---
 theories/common.v | 85 +++++++++++++++++++++++------------------------
 1 file changed, 41 insertions(+), 44 deletions(-)

diff --git a/theories/common.v b/theories/common.v
index b4867d2..76964a3 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -255,50 +255,39 @@ Inductive algo_dom : PTm -> PTm -> Prop :=
 | A_NatCong :
   algo_dom PNat PNat
 
-| A_NeuNeu a b :
-  algo_dom_neu a b ->
-  algo_dom a b
-
-| A_Conf a b :
-  ishf a ->
-  ishf b ->
-  tm_conf a b ->
-  algo_dom a b
-
-with algo_dom_neu : PTm -> PTm -> Prop :=
 | A_VarCong i j :
   (* -------------------------- *)
-  algo_dom_neu (VarPTm i) (VarPTm j)
+  algo_dom (VarPTm i) (VarPTm j)
 
 | A_AppCong u0 u1 a0 a1 :
   ishne u0 ->
   ishne u1 ->
-  algo_dom_neu u0 u1 ->
+  algo_dom u0 u1 ->
   algo_dom_r a0 a1 ->
   (* ------------------------- *)
-  algo_dom_neu (PApp u0 a0) (PApp u1 a1)
+  algo_dom (PApp u0 a0) (PApp u1 a1)
 
 | A_ProjCong p0 p1 u0 u1 :
   ishne u0 ->
   ishne u1 ->
-  algo_dom_neu u0 u1 ->
+  algo_dom u0 u1 ->
   (* ---------------------  *)
-  algo_dom_neu (PProj p0 u0) (PProj p1 u1)
+  algo_dom (PProj p0 u0) (PProj p1 u1)
 
 | A_IndCong P0 P1 u0 u1 b0 b1 c0 c1 :
   ishne u0 ->
   ishne u1 ->
   algo_dom_r P0 P1 ->
-  algo_dom_neu u0 u1 ->
+  algo_dom u0 u1 ->
   algo_dom_r b0 b1 ->
   algo_dom_r c0 c1 ->
-  algo_dom_neu (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
+  algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
 
-| A_NeuConf a b :
-  ishne a ->
-  ishne b ->
+| A_Conf a b :
+  HRed.nf a ->
+  HRed.nf b ->
   tm_conf a b ->
-  algo_dom_neu a b
+  algo_dom a b
 
 with algo_dom_r : PTm  -> PTm -> Prop :=
 | A_NfNf a b :
@@ -319,11 +308,10 @@ with algo_dom_r : PTm  -> PTm -> Prop :=
   algo_dom_r a b.
 
 Scheme algo_ind := Induction for algo_dom Sort Prop
-  with algo_neu_ind := Induction for algo_dom_neu Sort Prop
   with algor_ind := Induction for algo_dom_r Sort Prop.
 
-Combined Scheme algo_dom_mutual from algo_ind, algo_neu_ind, algor_ind.
-#[export]Hint Constructors algo_dom algo_dom_neu algo_dom_r : adom.
+Combined Scheme algo_dom_mutual from algo_ind, algor_ind.
+#[export]Hint Constructors algo_dom algo_dom_r : adom.
 
 Definition stm_nonconf a b :=
   match a, b with
@@ -332,9 +320,18 @@ Definition stm_nonconf a b :=
   | PBind PSig _ _, PBind PSig _ _ => true
   | PNat, PNat => true
   | VarPTm _, VarPTm _ => true
-  | PApp _ _, PApp _ _ => (~~ ishf a) && (~~ ishf b)
-  | PProj _ _, PProj _ _ => (~~ ishf a) && (~~ ishf b)
-  | PInd _ _ _ _, PInd _ _ _ _ => (~~ ishf a) && (~~ ishf b)
+  | PApp _ _, PApp _ _ => true
+  | PProj _ _, PProj _ _ => true
+  | PInd _ _ _ _, PInd _ _ _ _ => true
+  | _, _ => false
+  end.
+
+Definition neuneu_nonconf a b :=
+  match a, b with
+  | VarPTm _, VarPTm _ => true
+  | PApp _ _, PApp _ _ => true
+  | PProj _ _, PProj _ _ => true
+  | PInd _ _ _ _, PInd _ _ _ _ => true
   | _, _ => false
   end.
 
@@ -374,7 +371,8 @@ Inductive salgo_dom : PTm -> PTm -> Prop :=
   salgo_dom PNat PNat
 
 | S_NeuNeu a b :
-  algo_dom_neu a b ->
+  neuneu_nonconf a b ->
+  algo_dom a b ->
   salgo_dom a b
 
 | S_Conf a b :
@@ -401,6 +399,8 @@ with salgo_dom_r : PTm -> PTm -> Prop :=
   (* ----------------------- *)
   salgo_dom_r a b.
 
+#[export]Hint Constructors salgo_dom salgo_dom_r : sdom.
+
 Lemma hf_no_hred (a b : PTm) :
   ishf a ->
   HRed.R a b ->
@@ -446,16 +446,11 @@ Proof. induction 2; sauto. Qed.
 Lemma tm_conf_sym a b : tm_conf a b = tm_conf b a.
 Proof. case : a; case : b => //=. Qed.
 
-Lemma algo_dom_neu_hne (a b : PTm) :
-  algo_dom_neu a b ->
-  ishne a /\ ishne b.
-Proof. inversion 1; subst => //=. 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 use:algo_dom_neu_hne lq:on unfold:HRed.nf.
+  induction 1 =>//=; try hauto inv:HRed.R use:hne_no_hred, hf_no_hred lq:on unfold:HRed.nf.
 Qed.
 
 
@@ -473,7 +468,6 @@ Qed.
 
 Lemma algo_dom_sym :
   (forall a b (h : algo_dom a b), algo_dom b a) /\
-    (forall a b, algo_dom_neu a b -> algo_dom_neu b a) /\
     (forall a b (h : algo_dom_r a b), algo_dom_r b a).
 Proof.
   apply algo_dom_mutual; try qauto use:tm_conf_sym,A_HRedR' db:adom.
@@ -506,7 +500,7 @@ Lemma salgo_dom_no_hred (a b : PTm) :
   salgo_dom a b ->
   HRed.nf a /\ HRed.nf b.
 Proof.
-  induction 1 =>//=; try hauto inv:HRed.R use:hne_no_hred, hf_no_hred, algo_dom_neu_hne lq:on unfold:HRed.nf.
+  induction 1 =>//=; try hauto inv:HRed.R use:hne_no_hred, hf_no_hred, algo_dom_no_hred lq:on unfold:HRed.nf.
 Qed.
 
 Lemma S_HRedR' a b b' :
@@ -521,17 +515,20 @@ Proof.
   hauto lq:on use:S_HRedsL, S_HRedsR.
 Qed.
 
+Ltac solve_conf := intros; split;
+      apply S_Conf; solve [destruct_salgo | sfirstorder ctrs:salgo_dom use:hne_no_hred, hf_no_hred].
+
+Ltac solve_basic := hauto q:on ctrs:salgo_dom, salgo_dom_r, algo_dom use:algo_dom_sym.
+
 Lemma algo_dom_salgo_dom :
   (forall a b, algo_dom a b -> salgo_dom a b /\ salgo_dom b a) /\
-    (forall a b, algo_dom_neu a b -> True) /\
     (forall a b, algo_dom_r a b -> salgo_dom_r a b /\ salgo_dom_r b a).
 Proof.
-  apply algo_dom_mutual => //=;
-                             try hauto lq:on ctrs:salgo_dom, algo_dom_neu, salgo_dom_r use:S_Conf, hne_no_hred, algo_dom_sym, tm_stm_conf, S_HRedR' inv:HRed.R unfold:HRed.nf  solve+:destruct_salgo.
-  - case;case; hauto lq:on ctrs:salgo_dom, algo_dom_neu, salgo_dom_r use:S_Conf, hne_no_hred, algo_dom_sym inv:HRed.R unfold:HRed.nf  solve+:destruct_salgo.
-  - move => a b ha hb /[dup] /tm_stm_conf h.
-    rewrite tm_conf_sym => /tm_stm_conf h0.
-    hauto l:on use:S_Conf inv:HRed.R unfold:HRed.nf.
+  apply algo_dom_mutual => //=; try first [solve_conf | solve_basic].
+  - case; case; hauto lq:on ctrs:salgo_dom use:algo_dom_sym inv:HRed.R unfold:HRed.nf.
+  - move => a b ha hb hc. split;
+      apply S_Conf; hauto l:on use:tm_conf_sym, tm_stm_conf.
+  - hauto lq:on ctrs:salgo_dom_r use:S_HRedR'.
 Qed.
 
 Fixpoint hred (a : PTm) : option (PTm) :=

From 4021d05d99301b8befd9ab2125f1e136155948b8 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 18:18:42 -0400
Subject: [PATCH 191/210] Update executable to use salgo_dom for subtyping

---
 theories/algorithmic.v |  10 --
 theories/common.v      |   9 ++
 theories/executable.v  | 342 ++++++++++++++---------------------------
 3 files changed, 128 insertions(+), 233 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 17868a0..cf64209 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1051,12 +1051,6 @@ Proof.
     exists (S i), a1. hauto lq:on ctrs:nsteps solve+:lia.
 Qed.
 
-Lemma hreds_nf_refl a b  :
-  HRed.nf a ->
-  rtc HRed.R a b ->
-  a = b.
-Proof. inversion 2; sfirstorder. Qed.
-
 Lemma lored_nsteps_app_cong k (a0 a1 b : PTm) :
   nsteps LoRed.R k a0 a1 ->
   ishne a0 ->
@@ -1228,10 +1222,6 @@ Proof.
   hauto lq:on rew:off use:ne_nf b:on solve+:lia.
 Qed.
 
-
-Lemma algo_dom_r_algo_dom : forall a b, HRed.nf a -> HRed.nf b -> algo_dom_r a b -> algo_dom a b.
-Proof. hauto l:on use:algo_dom_r_inv, hreds_nf_refl. Qed.
-
 Lemma algo_dom_algo_dom_neu : forall a b, ishne a -> ishne b -> algo_dom a b -> algo_dom_neu a b \/ tm_conf a b.
 Proof.
   inversion 3; subst => //=; tauto.
diff --git a/theories/common.v b/theories/common.v
index 76964a3..266d90a 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -586,3 +586,12 @@ Definition fancy_hred (a : PTm) : HRed.nf a + {b | HRed.R a b}.
   right. exists p. by apply hred_sound in eq.
   left. move => b /hred_complete. congruence.
 Defined.
+
+Lemma hreds_nf_refl a b  :
+  HRed.nf a ->
+  rtc HRed.R a b ->
+  a = b.
+Proof. inversion 2; sfirstorder. Qed.
+
+Lemma algo_dom_r_algo_dom : forall a b, HRed.nf a -> HRed.nf b -> algo_dom_r a b -> algo_dom a b.
+Proof. hauto l:on use:algo_dom_r_inv, hreds_nf_refl. Qed.
diff --git a/theories/executable.v b/theories/executable.v
index fd03f2d..238fe45 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -11,133 +11,6 @@ Set Default Proof Mode "Classic".
 Require Import ssreflect ssrbool.
 From Hammer Require Import Tactics.
 
-Inductive eq_view : PTm -> PTm -> Type :=
-| V_AbsAbs a b :
-  eq_view (PAbs a) (PAbs b)
-| V_AbsNeu a b :
-  ~~ ishf b ->
-  eq_view (PAbs a) b
-| V_NeuAbs a b :
-  ~~ 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 :
-  ~~ ishf u ->
-  eq_view (PPair a0 b0) u
-| V_NeuPair u a1 b1 :
-  ~~ ishf 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 :
-  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) (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 (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;
-  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 salgo_dom : PTm -> PTm -> Prop := *)
-(* | S_UnivCong i j : *)
-(*   (* -------------------------- *) *)
-(*   salgo_dom (PUniv i) (PUniv j) *)
-
-(* | S_PiCong A0 A1 B0 B1 : *)
-(*   salgo_dom_r A1 A0 -> *)
-(*   salgo_dom_r B0 B1 -> *)
-(*   (* ---------------------------- *) *)
-(*   salgo_dom (PBind PPi A0 B0) (PBind PPi A1 B1) *)
-
-(* | S_SigCong A0 A1 B0 B1 : *)
-(*   salgo_dom_r A0 A1 -> *)
-(*   salgo_dom_r B0 B1 -> *)
-(*   (* ---------------------------- *) *)
-(*   salgo_dom (PBind PSig A0 B0) (PBind PSig A1 B1) *)
-
-(* | S_NatCong : *)
-(*   salgo_dom PNat PNat *)
-
-(* | S_NeuNeu a b : *)
-(*   ishne a -> *)
-(*   ishne b -> *)
-(*   algo_dom a b -> *)
-(*   (* ------------------- *) *)
-(*   salgo_dom *)
-
-(* | S_Conf a b : *)
-(*   HRed.nf a -> *)
-(*   HRed.nf b -> *)
-(*   tm_conf a b -> *)
-(*   salgo_dom a b *)
-
-(* with salgo_dom_r : PTm  -> PTm -> Prop := *)
-(* | S_NfNf a b : *)
-(*   salgo_dom a b -> *)
-(*   salgo_dom_r a b *)
-
-(* | S_HRedL a a' b  : *)
-(*   HRed.R a a' -> *)
-(*   salgo_dom_r a' b -> *)
-(*   (* ----------------------- *) *)
-(*   salgo_dom_r a b *)
-
-(* | S_HRedR a b b'  : *)
-(*   HRed.nf a -> *)
-(*   HRed.R b b' -> *)
-(*   salgo_dom_r a b' -> *)
-(*   (* ----------------------- *) *)
-(*   salgo_dom_r a b. *)
-
-(* Scheme salgo_ind := Induction for salgo_dom Sort Prop *)
-(*   with algor_ind := Induction for salgo_dom_r Sort Prop. *)
-
-
 Ltac2 destruct_algo () :=
   lazy_match! goal with
   | [h : algo_dom ?a ?b |- _ ] =>
@@ -161,70 +34,79 @@ Ltac solve_check_equal :=
   | _ => idtac
   end].
 
-#[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;
-  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
+Global Set Transparent Obligations.
 
-  (* 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 _.
+Local Obligation Tactic := try solve [check_equal_triv | sfirstorder].
+
+Program Fixpoint check_equal (a b : PTm) (h : algo_dom a b) {struct h} : bool :=
+  match a, b with
+  | VarPTm i, VarPTm j => nat_eqdec i j
+  | PAbs a, PAbs b => check_equal_r a b _
+  | PAbs a, VarPTm _ => check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) _
+  | PAbs a, PApp _ _ => check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) _
+  | PAbs a, PInd _ _ _ _ => check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) _
+  | PAbs a, PProj _ _ => check_equal_r a (PApp (ren_PTm shift b) (VarPTm var_zero)) _
+  | VarPTm _, PAbs b => check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b _
+  | PApp _ _, PAbs b => check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b _
+  | PProj _ _, PAbs b => check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b _
+  | PInd _ _ _ _, PAbs b => check_equal_r (PApp (ren_PTm shift a) (VarPTm var_zero)) b _
+  | PPair a0 b0, PPair a1 b1 =>
+      check_equal_r a0 a1 _ && check_equal_r b0 b1 _
+  | PPair a0 b0, VarPTm _ => check_equal_r a0 (PProj PL b) _ && check_equal_r b0 (PProj PR b) _
+  | PPair a0 b0, PProj _ _ => check_equal_r a0 (PProj PL b) _ && check_equal_r b0 (PProj PR b) _
+  | PPair a0 b0, PApp _ _ => check_equal_r a0 (PProj PL b) _ && check_equal_r b0 (PProj PR b) _
+  | PPair a0 b0, PInd _ _ _ _ => check_equal_r a0 (PProj PL b) _ && check_equal_r b0 (PProj PR b) _
+  | VarPTm _, PPair a1 b1 => check_equal_r (PProj PL a) a1 _ && check_equal_r (PProj PR a) b1 _
+  | PApp _ _, PPair a1 b1 => check_equal_r (PProj PL a) a1 _ && check_equal_r (PProj PR a) b1 _
+  | PProj _ _, PPair a1 b1 => check_equal_r (PProj PL a) a1 _ && check_equal_r (PProj PR a) b1 _
+  | PInd _ _ _ _, PPair a1 b1 => check_equal_r (PProj PL a) a1 _ && check_equal_r (PProj PR a) b1 _
+  | PNat, PNat => true
+  | PZero, PZero => true
+  | PSuc a, PSuc b => check_equal_r a b _
+  | PProj p0 a, PProj p1 b => PTag_eqdec p0 p1 && check_equal a b _
+  | PApp a0 b0, PApp a1 b1 => check_equal a0 a1 _ && check_equal_r b0 b1 _
+  | PInd P0 u0 b0 c0, PInd P1 u1 b1 c1 =>
+      check_equal_r P0 P1 _ && check_equal u0 u1 _ && check_equal_r b0 b1 _ && check_equal_r c0 c1 _
+  | PUniv i, PUniv j => nat_eqdec i j
+  | PBind p0 A0 B0, PBind p1 A1 B1 => BTag_eqdec p0 p1 && check_equal_r A0 A1 _ && check_equal_r B0 B1 _
+  | _, _ => false
+  end
+with check_equal_r (a b : PTm) (h : algo_dom_r a b) {struct h} : bool :=
+  match fancy_hred a with
+  | inr a' => check_equal_r (proj1_sig a') b _
+  | inl ha' => match fancy_hred b with
+            | inr b' => check_equal_r a (proj1_sig b') _
+            | inl hb' => check_equal a b _
+            end
+  end.
 
 Next Obligation.
-  intros.
+  simpl. intros. clear Heq_anonymous.  destruct a' as [a' ha']. simpl.
   inversion h; subst => //=.
-  exfalso. hauto l:on use:hred_complete unfold:HRed.nf.
-  exfalso. hauto l:on use:hred_complete unfold:HRed.nf.
-Defined.
-
-Next Obligation.
-  intros.
-  destruct h.
-  exfalso. sfirstorder use: algo_dom_no_hred.
-  exfalso. sfirstorder.
-  assert (  b' = b'0)by eauto using hred_deter. subst.
-  apply h.
-Defined.
-
-Next Obligation.
-  simpl. intros.
-  inversion h; subst =>//=.
   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.
+  assert (a' = a'0) by eauto using hred_deter. by subst.
+  exfalso. sfirstorder.
 Defined.
 
+Next Obligation.
+  simpl. intros. clear Heq_anonymous Heq_anonymous0.
+  destruct b' as [b' hb']. simpl.
+  inversion h; subst.
+  - exfalso.
+    sfirstorder use:algo_dom_no_hred.
+  - exfalso.
+    sfirstorder.
+  - assert (b' = b'0) by eauto using hred_deter. by subst.
+Defined.
+
+(* Need to avoid ssreflect tactics since they generate terms that make the termination checker upset *)
+Next Obligation.
+  move => /= a b hdom ha _ hb _.
+  inversion hdom; subst.
+  - assumption.
+  - exfalso; sfirstorder.
+  - exfalso; sfirstorder.
+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.
@@ -279,14 +161,14 @@ Proof.
   sfirstorder use:hred_complete, hred_sound.
 Qed.
 
+Ltac simp_check_r := with_strategy opaque [check_equal] simpl in *.
+
 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.
+  simp_check_r.
   destruct (fancy_hred a).
-  simpl.
   destruct (fancy_hred b).
   reflexivity.
   exfalso. hauto l:on use:hred_complete.
@@ -297,11 +179,9 @@ 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.
@@ -312,7 +192,7 @@ 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.
+  simpl.
   destruct (fancy_hred a).
   - simpl.
     destruct (fancy_hred b) as [|[b'' hb']].
@@ -335,31 +215,51 @@ 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.
 
-Obligation Tactic := try solve [check_equal_triv | sfirstorder].
+Ltac2 destruct_salgo () :=
+  lazy_match! goal with
+  | [h : salgo_dom ?a ?b |- _ ] =>
+      if is_var a then destruct $a; ltac1:(done)  else
+        (if is_var b then destruct $b; ltac1:(done) else ())
+  end.
 
-Program Fixpoint check_sub (q : bool) (a b : PTm) (h : algo_dom a b) {struct h} :=
+Ltac check_sub_triv :=
+  intros;subst;
+  lazymatch goal with
+  (* | [h : algo_dom (VarPTm _) (PAbs _) |- _] => idtac *)
+  | [h : salgo_dom _ _ |- _] => try (inversion h; subst => //=; ltac2:(Control.enter destruct_algo))
+  | _ => idtac
+  end.
+
+Local Obligation Tactic := try solve [check_sub_triv | sfirstorder].
+
+Program Fixpoint check_sub (a b : PTm) (h : salgo_dom a b) {struct h} :=
   match a, b with
   | PBind PPi A0 B0, PBind PPi A1 B1 =>
-      check_sub_r (negb q) A0 A1 _ && check_sub_r q B0 B1 _
+      check_sub_r A1 A0 _ && check_sub_r B0 B1 _
   | PBind PSig A0 B0, PBind PSig A1 B1 =>
-      check_sub_r q A0 A1 _ && check_sub_r q B0 B1 _
+      check_sub_r A0 A1 _ && check_sub_r B0 B1 _
   | PUniv i, PUniv j =>
-      if q then PeanoNat.Nat.leb i j else PeanoNat.Nat.leb j i
+      PeanoNat.Nat.leb i j
   | PNat, PNat => true
-  | _ ,_ => ishne a && ishne b && check_equal a b h
+  | PApp _ _ , PApp _ _ => check_equal a b _
+  | VarPTm _, VarPTm _ => check_equal a b _
+  | PInd _ _ _ _, PInd _ _ _ _ => check_equal a b _
+  | PProj _ _, PProj _ _ => check_equal a b _
+  | _, _ => false
   end
-with check_sub_r (q : bool) (a b : PTm) (h : algo_dom_r a b) {struct h} :=
+with check_sub_r (a b : PTm) (h : salgo_dom_r a b) {struct h} :=
   match fancy_hred a with
-  | inr a' => check_sub_r q (proj1_sig a') b _
+  | inr a' => check_sub_r (proj1_sig a') b _
   | inl ha' => match fancy_hred b with
-            | inr b' => check_sub_r q a (proj1_sig b') _
-            | inl hb' => check_sub q a b _
+            | inr b' => check_sub_r a (proj1_sig b') _
+            | inl hb' => check_sub  a b _
             end
   end.
+
 Next Obligation.
   simpl. intros. clear Heq_anonymous.  destruct a' as [a' ha']. simpl.
   inversion h; subst => //=.
-  exfalso. sfirstorder use:algo_dom_no_hred.
+  exfalso. sfirstorder use:salgo_dom_no_hred.
   assert (a' = a'0) by eauto using hred_deter. by subst.
   exfalso. sfirstorder.
 Defined.
@@ -369,7 +269,7 @@ Next Obligation.
   destruct b' as [b' hb']. simpl.
   inversion h; subst.
   - exfalso.
-    sfirstorder use:algo_dom_no_hred.
+    sfirstorder use:salgo_dom_no_hred.
   - exfalso.
     sfirstorder.
   - assert (b' = b'0) by eauto using hred_deter. by subst.
@@ -377,34 +277,30 @@ Defined.
 
 (* Need to avoid ssreflect tactics since they generate terms that make the termination checker upset *)
 Next Obligation.
-  move => _ /= a b hdom ha _ hb _.
+  move => /= a b hdom ha _ hb _.
   inversion hdom; subst.
   - assumption.
   - exfalso; sfirstorder.
   - exfalso; sfirstorder.
 Defined.
 
-Lemma check_sub_pi_pi q A0 B0  A1 B1 h0 h1 :
-  check_sub q (PBind PPi A0 B0) (PBind PPi A1 B1) (A_BindCong PPi PPi A0 A1 B0 B1 h0 h1) =
-    check_sub_r (~~ q) A0 A1 h0 && check_sub_r q B0 B1 h1.
+Lemma check_sub_pi_pi A0 B0  A1 B1 h0 h1 :
+  check_sub (PBind PPi A0 B0) (PBind PPi A1 B1) (S_PiCong A0 A1 B0 B1 h0 h1) =
+    check_sub_r A1 A0 h0 && check_sub_r B0 B1 h1.
 Proof. reflexivity. Qed.
 
-Lemma check_sub_sig_sig q A0 B0  A1 B1 h0 h1 :
-  check_sub q (PBind PSig A0 B0) (PBind PSig A1 B1) (A_BindCong PSig PSig A0 A1 B0 B1 h0 h1) =
-    check_sub_r q A0 A1 h0 && check_sub_r q B0 B1 h1.
-  reflexivity. Qed.
+Lemma check_sub_sig_sig A0 B0  A1 B1 h0 h1 :
+  check_sub (PBind PSig A0 B0) (PBind PSig A1 B1) (S_SigCong A0 A1 B0 B1 h0 h1) =
+    check_sub_r A0 A1 h0 && check_sub_r B0 B1 h1.
+Proof. reflexivity. Qed.
 
 Lemma check_sub_univ_univ i j :
-  check_sub true (PUniv i) (PUniv j) (A_UnivCong i j) = PeanoNat.Nat.leb i j.
+  check_sub (PUniv i) (PUniv j) (S_UnivCong i j) = PeanoNat.Nat.leb i j.
 Proof. reflexivity. Qed.
 
-Lemma check_sub_univ_univ' i j :
-  check_sub false (PUniv i) (PUniv j) (A_UnivCong i j) = PeanoNat.Nat.leb j i.
-  reflexivity. Qed.
-
-Lemma check_sub_nfnf q a b dom : check_sub_r q a b (A_NfNf a b dom) = check_sub q a b dom.
+Lemma check_sub_nfnf a b dom : check_sub_r a b (S_NfNf a b dom) = check_sub a b dom.
 Proof.
-  have [h0 h1] : HRed.nf a /\ HRed.nf b by hauto l:on use:algo_dom_no_hred.
+  have [h0 h1] : HRed.nf a /\ HRed.nf b by hauto l:on use:salgo_dom_no_hred.
   have [h3 h4] : hred a = None /\ hred b = None by sfirstorder use:hf_no_hred, hne_no_hred, hred_none.
   simpl.
   destruct (fancy_hred b)=>//=.
@@ -415,8 +311,8 @@ Proof.
   hauto l:on use:hred_complete.
 Qed.
 
-Lemma check_sub_hredl q a b a' ha doma :
-  check_sub_r q a b (A_HRedL a a' b ha doma) = check_sub_r q a' b doma.
+Lemma check_sub_hredl a b a' ha doma :
+  check_sub_r a b (S_HRedL a a' b ha doma) = check_sub_r a' b doma.
 Proof.
   simpl.
   destruct (fancy_hred a).
@@ -428,9 +324,9 @@ Proof.
     apply PropExtensionality.proof_irrelevance.
 Qed.
 
-Lemma check_sub_hredr q a b b' hu r a0 :
-  check_sub_r q a b (A_HRedR a b b' hu r a0) =
-    check_sub_r q a b' a0.
+Lemma check_sub_hredr a b b' hu r a0 :
+  check_sub_r a b (S_HRedR a b b' hu r a0) =
+    check_sub_r a b' a0.
 Proof.
   simpl.
   destruct (fancy_hred a).
@@ -445,4 +341,4 @@ Proof.
     sfirstorder use:hne_no_hred, hf_no_hred.
 Qed.
 
-#[export]Hint Rewrite check_sub_hredl check_sub_hredr check_sub_nfnf check_sub_univ_univ' check_sub_univ_univ check_sub_pi_pi check_sub_sig_sig : ce_prop.
+#[export]Hint Rewrite check_sub_hredl check_sub_hredr check_sub_nfnf check_sub_univ_univ check_sub_pi_pi check_sub_sig_sig : ce_prop.

From 030dccb3261e8ee8d3bfde3e4a455b85b936c0c9 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 19:02:51 -0400
Subject: [PATCH 192/210] Finish the refactored executable_correct

---
 theories/algorithmic.v        | 1522 ++++++++++++++++-----------------
 theories/common.v             |    2 +
 theories/executable_correct.v |  203 ++---
 3 files changed, 847 insertions(+), 880 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index cf64209..de58fd4 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -739,7 +739,7 @@ Proof.
   move => ha hb.
   move => ?.
   apply A_NfNf.
-  by apply A_Conf.
+  apply A_Conf; sfirstorder use:hf_no_hred, hne_no_hred.
 Qed.
 
 Lemma hne_nf_ne (a : PTm ) :
@@ -1222,11 +1222,6 @@ Proof.
   hauto lq:on rew:off use:ne_nf b:on solve+:lia.
 Qed.
 
-Lemma algo_dom_algo_dom_neu : forall a b, ishne a -> ishne b -> algo_dom a b -> algo_dom_neu a b \/ tm_conf a b.
-Proof.
-  inversion 3; subst => //=; tauto.
-Qed.
-
 Lemma term_metric_algo_dom : forall k a b, term_metric k a b -> algo_dom_r a b.
 Proof.
   move => [:hneL].
@@ -1266,8 +1261,7 @@ Proof.
       have nfa0' : HRed.nf a0 by sfirstorder use:hne_no_hred.
       have nfb0' : HRed.nf a1 by sfirstorder use:hne_no_hred.
       have ha0 : algo_dom a0 a1 by eauto using algo_dom_r_algo_dom.
-      apply A_Conf. admit. admit. rewrite /tm_conf. simpl.
-      eauto using algo_dom_r_algo_dom.
+      constructor => //. eauto.
     + move => p0 A0 p1 A1 neA0 neA1.
       have {}nfa0 : HRed.nf A0 by sfirstorder use:hne_no_hred.
       have {}nfb0 : HRed.nf A1 by sfirstorder use:hne_no_hred.
@@ -1278,489 +1272,489 @@ Proof.
       hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
 Qed.
 
-Lemma coqeq_complete' k (a b : PTm ) :
-  (forall a b, algo_dom a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b)
+(* Lemma coqeq_complete' k (a b : PTm ) : *)
+(*   (forall a b, algo_dom a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) *)
 
-  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 a b.
-  elim /Wf_nat.lt_wf_ind.
-  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 : 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 /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 : a b h. move => [:hnfneu].
-  move => a b h.
-  case => fb; case => fa.
-  - split; last by sfirstorder use:hf_not_hne.
-    move {hnfneu}.
-    case : a h fb fa => //=.
-    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp.
-      move => a0 a1 ha0 _ _ Γ A wt0 wt1.
-      move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]].
-      apply : CE_HRed; eauto using rtc_refl.
-      apply CE_AbsAbs.
-      suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto.
-      have ? : n > 0 by sauto solve+:lia.
-      exists (n - 1). split; first by lia.
-      move : (ha0). rewrite /algo_metric.
-      move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har.
-      apply lored_nsteps_abs_inv in hr0, hr1.
-      move : hr0 => [va' [hr00 hr01]].
-      move : hr1 => [vb' [hr10 hr11]]. move {ih}.
-      exists i,j,va',vb'. subst.
-      suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
-      repeat split =>//=. sfirstorder.
-      simpl in *; by lia.
-      move /algo_metric_join /DJoin.symmetric : ha0.
-      have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
-      move => /[dup] [[ha00 ha10]] [].
-      move : DJoin.abs_inj; repeat move/[apply].
-      move : DJoin.standardization ha00 ha10; repeat move/[apply].
-      (* this is awful *)
-      move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]].
-      have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va'
-        by hauto lq:on use:@relations.rtc_nsteps.
-      have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb'
-        by hauto lq:on use:@relations.rtc_nsteps.
-      simpl in *.
-      have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst.
-      sfirstorder.
-    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp.
-      move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
-      have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1)
-        by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
-      apply : CE_HRed; eauto using rtc_refl.
-      move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0.
-      move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1.
-      move /Sub_Bind_InvR :  (hSu0).
-      move => [i][A2][B2]hE.
-      have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2
-        by eauto using Su_Transitive, Su_Eq.
-      have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2
-        by eauto using Su_Transitive, Su_Eq.
-      have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1.
-      have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1.
-      have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
-      have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
-      move /algo_metric_pair : h sn0 sn1; repeat move/[apply].
-      move => [j][hj][ih0 ih1].
-      have haE : a0 ⇔ a1 by hauto l:on.
-      apply CE_PairPair => //.
-      have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2
-        by hauto l:on use:coqeq_sound_mutual.
-      have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2.
-      apply : T_Conv; eauto.
-      move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2.
-      eapply ih; cycle -1; eauto.
-      apply : T_Conv; eauto.
-      apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2).
-      move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2.
-      move : hE haE. clear.
-      move => h.
-      eapply regularity in h.
-      move : h => [_ [hB _]].
-      eauto using bind_inst.
-    + case : b => //=.
-      * hauto lq:on use:T_AbsBind_Imp.
-      * hauto lq:on rew:off use:T_PairBind_Imp.
-      * move => p1 A1 B1 p0 A0 B0.
-        move /algo_metric_bind.
-        move => [?][j [ih0 [ih1 ih2]]]_ _. subst.
-        move => Γ A hU0 hU1.
-        move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]].
-        move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]].
-        have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1)
-          by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia.
-        have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1)
-          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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1)
-          by hauto lq:on use:T_Univ_Raise solve+:lia.
-        have {}hB1 : (cons A1 Γ) ⊢ B1 ∈ PUniv (max i0 i1)
-          by hauto lq:on use:T_Univ_Raise solve+:lia.
+(*   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 a b. *)
+(*   elim /Wf_nat.lt_wf_ind. *)
+(*   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 : 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 /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 : a b h. move => [:hnfneu]. *)
+(*   move => a b h. *)
+(*   case => fb; case => fa. *)
+(*   - split; last by sfirstorder use:hf_not_hne. *)
+(*     move {hnfneu}. *)
+(*     case : a h fb fa => //=. *)
+(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp. *)
+(*       move => a0 a1 ha0 _ _ Γ A wt0 wt1. *)
+(*       move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]]. *)
+(*       apply : CE_HRed; eauto using rtc_refl. *)
+(*       apply CE_AbsAbs. *)
+(*       suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto. *)
+(*       have ? : n > 0 by sauto solve+:lia. *)
+(*       exists (n - 1). split; first by lia. *)
+(*       move : (ha0). rewrite /algo_metric. *)
+(*       move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har. *)
+(*       apply lored_nsteps_abs_inv in hr0, hr1. *)
+(*       move : hr0 => [va' [hr00 hr01]]. *)
+(*       move : hr1 => [vb' [hr10 hr11]]. move {ih}. *)
+(*       exists i,j,va',vb'. subst. *)
+(*       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0. *)
+(*       repeat split =>//=. sfirstorder. *)
+(*       simpl in *; by lia. *)
+(*       move /algo_metric_join /DJoin.symmetric : ha0. *)
+(*       have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
+(*       move => /[dup] [[ha00 ha10]] []. *)
+(*       move : DJoin.abs_inj; repeat move/[apply]. *)
+(*       move : DJoin.standardization ha00 ha10; repeat move/[apply]. *)
+(*       (* this is awful *) *)
+(*       move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]]. *)
+(*       have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va' *)
+(*         by hauto lq:on use:@relations.rtc_nsteps. *)
+(*       have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb' *)
+(*         by hauto lq:on use:@relations.rtc_nsteps. *)
+(*       simpl in *. *)
+(*       have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst. *)
+(*       sfirstorder. *)
+(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp. *)
+(*       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1. *)
+(*       have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1) *)
+(*         by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
+(*       apply : CE_HRed; eauto using rtc_refl. *)
+(*       move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0. *)
+(*       move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1. *)
+(*       move /Sub_Bind_InvR :  (hSu0). *)
+(*       move => [i][A2][B2]hE. *)
+(*       have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2 *)
+(*         by eauto using Su_Transitive, Su_Eq. *)
+(*       have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 *)
+(*         by eauto using Su_Transitive, Su_Eq. *)
+(*       have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1. *)
+(*       have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1. *)
+(*       have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
+(*       have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
+(*       move /algo_metric_pair : h sn0 sn1; repeat move/[apply]. *)
+(*       move => [j][hj][ih0 ih1]. *)
+(*       have haE : a0 ⇔ a1 by hauto l:on. *)
+(*       apply CE_PairPair => //. *)
+(*       have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2 *)
+(*         by hauto l:on use:coqeq_sound_mutual. *)
+(*       have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2. *)
+(*       apply : T_Conv; eauto. *)
+(*       move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2. *)
+(*       eapply ih; cycle -1; eauto. *)
+(*       apply : T_Conv; eauto. *)
+(*       apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2). *)
+(*       move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2. *)
+(*       move : hE haE. clear. *)
+(*       move => h. *)
+(*       eapply regularity in h. *)
+(*       move : h => [_ [hB _]]. *)
+(*       eauto using bind_inst. *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on use:T_AbsBind_Imp. *)
+(*       * hauto lq:on rew:off use:T_PairBind_Imp. *)
+(*       * move => p1 A1 B1 p0 A0 B0. *)
+(*         move /algo_metric_bind. *)
+(*         move => [?][j [ih0 [ih1 ih2]]]_ _. subst. *)
+(*         move => Γ A hU0 hU1. *)
+(*         move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]]. *)
+(*         move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]]. *)
+(*         have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1) *)
+(*           by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia. *)
+(*         have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1) *)
+(*           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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1) *)
+(*           by hauto lq:on use:T_Univ_Raise solve+:lia. *)
+(*         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.
-        have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1)
-          by hauto l:on use:coqeq_sound_mutual.
-        apply : CE_HRed; eauto using rtc_refl.
-        constructor => //.
+(*         have ihA : A0 ⇔ A1 by hauto l:on. *)
+(*         have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1) *)
+(*           by hauto l:on use:coqeq_sound_mutual. *)
+(*         apply : CE_HRed; eauto using rtc_refl. *)
+(*         constructor => //. *)
 
-        eapply ih; eauto.
-        apply : ctx_eq_subst_one; eauto.
-        eauto using Su_Eq.
-      * move => > /algo_metric_join.
-        hauto lq:on use:DJoin.bind_univ_noconf.
-      * move => > /algo_metric_join.
-        hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin.
-      * move => > /algo_metric_join.
-        clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv.
-      * move => > /algo_metric_join. clear.
-        hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv.
-    + case : b => //=.
-      * hauto lq:on use:T_AbsUniv_Imp.
-      * hauto lq:on use:T_PairUniv_Imp.
-      * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric.
-      * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj.
-        subst.
-        hauto l:on.
-      * move => > /algo_metric_join.
-        hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin.
-      * move => > /algo_metric_join.
-        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv.
-      * move => > /algo_metric_join.
-        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv.
-    + case : b => //=.
-      * qauto l:on use:T_AbsNat_Imp.
-      * qauto l:on use:T_PairNat_Imp.
-      * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf.
-      * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf.
-      * hauto l:on.
-      * move /algo_metric_join.
-        hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv.
-    (* Zero *)
-    + case : b => //=.
-      * hauto lq:on rew:off use:T_AbsZero_Imp.
-      * hauto lq: on use: T_PairZero_Imp.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv.
-      * hauto l:on.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv.
-    (* Suc *)
-    + case : b => //=.
-      * hauto lq:on rew:off use:T_AbsSuc_Imp.
-      * hauto lq:on use:T_PairSuc_Imp.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv.
-      * move => a0 a1 h _ _.
-        move /algo_metric_suc : h => [j [h4 h5]].
-        move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3].
-        move : ih h4 h5;do!move/[apply].
-        move => [ih _].
-        move : ih h0 h2;do!move/[apply].
-        move => h. apply : CE_HRed; eauto using rtc_refl.
-        by constructor.
-  - move : a b h fb fa. abstract : hnfneu.
-    move => + b.
-    case : b => //=.
-    (* NeuAbs *)
-    + move => a u halg _ nea. split => // Γ A hu /[dup] ha.
-      move /Abs_Inv => [A0][B0][_]hSu.
-      move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE.
-      have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
-      have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
-      have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i
-        by hauto l:on use:regularity.
-      have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j
-          by qauto l:on use:Bind_Inv.
-      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.
-      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.
-      apply : CE_HRed; eauto using rtc_refl.
-      eapply CE_NeuAbs => //=.
-      eapply ih; eauto.
-    (* NeuPair *)
-    + move => a0 b0 u halg _ neu.
-      split => // Γ A hu /[dup] wt.
-      move /Pair_Inv => [A0][B0][ha0][hb0]hU.
-      move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE.
-      have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on.
-      have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
-      move /Pair_Sig_Inv : wt => [{}ha0 {}hb0].
-      have /T_Proj1 huL := hu.
-      have /T_Proj2 {hu} huR := hu.
-      move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]].
-      have heL : PProj PL u ⇔ a0 by hauto l:on.
-      eapply CE_HRed; eauto using rtc_refl.
-      apply CE_NeuPair; eauto.
-      eapply ih; eauto.
-      apply : T_Conv; eauto.
-      have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i
-        by hauto l:on use:regularity.
-      have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by
-        hauto l:on use:coqeq_sound_mutual.
-      hauto l:on use:bind_inst.
-    (* NeuBind: Impossible *)
-    + move => b p p0 a /algo_metric_join h _ h0. exfalso.
-      move : h h0. clear.
-      move /Sub.FromJoin.
-      hauto l:on use:Sub.hne_bind_noconf.
-    (* NeuUniv: Impossible *)
-    + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join.
-    + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join.
-    (* Zero *)
-    + case => //=.
-      * move => i /algo_metric_join. clear.
-        hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv.
-      * move => >/algo_metric_join. clear.
-        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.
-      * move => >/algo_metric_join. clear.
-        move => h _ h2. exfalso.
-        hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv.
-    (* Suc *)
-    + move => a0.
-      case => //=; move => >/algo_metric_join; clear.
-      * 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.
-      * 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 Γ 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.
+(*         eapply ih; eauto. *)
+(*         apply : ctx_eq_subst_one; eauto. *)
+(*         eauto using Su_Eq. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on use:DJoin.bind_univ_noconf. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin. *)
+(*       * move => > /algo_metric_join. *)
+(*         clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv. *)
+(*       * move => > /algo_metric_join. clear. *)
+(*         hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv. *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on use:T_AbsUniv_Imp. *)
+(*       * hauto lq:on use:T_PairUniv_Imp. *)
+(*       * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric. *)
+(*       * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj. *)
+(*         subst. *)
+(*         hauto l:on. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin. *)
+(*       * move => > /algo_metric_join. *)
+(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv. *)
+(*     + case : b => //=. *)
+(*       * qauto l:on use:T_AbsNat_Imp. *)
+(*       * qauto l:on use:T_PairNat_Imp. *)
+(*       * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf. *)
+(*       * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf. *)
+(*       * hauto l:on. *)
+(*       * move /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv. *)
+(*     (* Zero *) *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on rew:off use:T_AbsZero_Imp. *)
+(*       * hauto lq: on use: T_PairZero_Imp. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv. *)
+(*       * hauto l:on. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv. *)
+(*     (* Suc *) *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on rew:off use:T_AbsSuc_Imp. *)
+(*       * hauto lq:on use:T_PairSuc_Imp. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv. *)
+(*       * move => a0 a1 h _ _. *)
+(*         move /algo_metric_suc : h => [j [h4 h5]]. *)
+(*         move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3]. *)
+(*         move : ih h4 h5;do!move/[apply]. *)
+(*         move => [ih _]. *)
+(*         move : ih h0 h2;do!move/[apply]. *)
+(*         move => h. apply : CE_HRed; eauto using rtc_refl. *)
+(*         by constructor. *)
+(*   - move : a b h fb fa. abstract : hnfneu. *)
+(*     move => + b. *)
+(*     case : b => //=. *)
+(*     (* NeuAbs *) *)
+(*     + move => a u halg _ nea. split => // Γ A hu /[dup] ha. *)
+(*       move /Abs_Inv => [A0][B0][_]hSu. *)
+(*       move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE. *)
+(*       have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
+(*       have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
+(*       have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i *)
+(*         by hauto l:on use:regularity. *)
+(*       have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j *)
+(*           by qauto l:on use:Bind_Inv. *)
+(*       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. *)
+(*       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. *)
+(*       apply : CE_HRed; eauto using rtc_refl. *)
+(*       eapply CE_NeuAbs => //=. *)
+(*       eapply ih; eauto. *)
+(*     (* NeuPair *) *)
+(*     + move => a0 b0 u halg _ neu. *)
+(*       split => // Γ A hu /[dup] wt. *)
+(*       move /Pair_Inv => [A0][B0][ha0][hb0]hU. *)
+(*       move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE. *)
+(*       have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on. *)
+(*       have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
+(*       move /Pair_Sig_Inv : wt => [{}ha0 {}hb0]. *)
+(*       have /T_Proj1 huL := hu. *)
+(*       have /T_Proj2 {hu} huR := hu. *)
+(*       move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]]. *)
+(*       have heL : PProj PL u ⇔ a0 by hauto l:on. *)
+(*       eapply CE_HRed; eauto using rtc_refl. *)
+(*       apply CE_NeuPair; eauto. *)
+(*       eapply ih; eauto. *)
+(*       apply : T_Conv; eauto. *)
+(*       have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i *)
+(*         by hauto l:on use:regularity. *)
+(*       have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by *)
+(*         hauto l:on use:coqeq_sound_mutual. *)
+(*       hauto l:on use:bind_inst. *)
+(*     (* NeuBind: Impossible *) *)
+(*     + move => b p p0 a /algo_metric_join h _ h0. exfalso. *)
+(*       move : h h0. clear. *)
+(*       move /Sub.FromJoin. *)
+(*       hauto l:on use:Sub.hne_bind_noconf. *)
+(*     (* NeuUniv: Impossible *) *)
+(*     + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join. *)
+(*     + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join. *)
+(*     (* Zero *) *)
+(*     + case => //=. *)
+(*       * move => i /algo_metric_join. clear. *)
+(*         hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv. *)
+(*       * move => >/algo_metric_join. clear. *)
+(*         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. *)
+(*       * move => >/algo_metric_join. clear. *)
+(*         move => h _ h2. exfalso. *)
+(*         hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv. *)
+(*     (* Suc *) *)
+(*     + move => a0. *)
+(*       case => //=; move => >/algo_metric_join; clear. *)
+(*       * 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. *)
+(*       * 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 Γ 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 => 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.
-      * clear => ? ? _. exfalso.
-        hauto q:on use:REReds.var_inv, REReds.hne_ind_inv.
-    + case : b => //=;
-                   lazymatch goal with
-                   | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac
-                   | _ => move => > /algo_metric_join
-                   end.
-      * clear => *. exfalso.
-        hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv.
-      (* real case *)
-      * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
-        move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
-        move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1.
-        move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply].
-        move => [j [hj [hal0 hal1]]].
-        have /ih := hal0.
-        move /(_ hj).
-        move => [_ ihb].
-        move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply].
-        move => [hb01][C][hT0][hT1][hT2]hT3.
-        move /Sub_Bind_InvL : (hT0).
-        move => [i][A2][B2]hE.
-        have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by
-          eauto using Su_Eq, Su_Transitive.
-        have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by
-          eauto using Su_Eq, Su_Transitive.
-        have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
-        have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
-        have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
-        have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
-        move : ih (hj) hal1. repeat move/[apply].
-        move => [ih _].
-        move : ih (ha0') (ha1'); repeat move/[apply].
-        move => iha.
-        split. qblast.
-        exists (subst_PTm (scons a0 VarPTm) B2).
-        split.
-        apply : Su_Transitive; eauto.
-        move /E_Refl in ha0.
-        hauto l:on use:Su_Pi_Proj2.
-        have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
-        split.
-        apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
-        move /regularity_sub0 : hSu10 => [i0].
-        hauto l:on use:bind_inst.
-        hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
-        split.
-        by apply : T_App; eauto using T_Conv_E.
-        apply : T_Conv; eauto.
-        apply T_App with (A := A2) (B := B2); eauto.
-        apply : T_Conv_E; eauto.
-        move /E_Symmetric in h01.
-        move /regularity_sub0 : hSu20 => [i0].
-        sfirstorder use:bind_inst.
-      * clear => ? ? ?. exfalso.
-        hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
-      * clear => ? ? ?. exfalso.
-        hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv.
-    + case : b => //=.
-      * move => i p h /algo_metric_join. clear => ? _ ?. exfalso.
-        hauto l:on use:REReds.hne_proj_inv, REReds.var_inv.
-      * move => > /algo_metric_join. clear => ? ? ?. exfalso.
-        hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv.
-      (* real case *)
-      * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0.
-        move : ha (hne0) (hne1); repeat move/[apply].
-        move => [? [j []]]. subst.
-        move : ih; repeat move/[apply].
-        move => [_ ih].
-        case : p1.
-        ** move => Γ A B ha0 ha1.
-           move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0.
-           move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1.
-           move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply].
-           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
-           split. sauto lq:on.
-           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
-           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
-             by eauto using Su_Transitive, Su_Eq.
-           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
-             by eauto using Su_Transitive, Su_Eq.
-           exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive.
-           repeat split => //=.
-           hauto l:on use:Su_Sig_Proj1, Su_Transitive.
-           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
-           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
-        ** move => Γ A B  ha0  ha1.
-           move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0.
-           move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1.
-           move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply].
-           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
-           split. sauto lq:on.
-           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
-           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
-             by eauto using Su_Transitive, Su_Eq.
-           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
-             by eauto using Su_Transitive, Su_Eq.
-           have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1.
-           have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1.
-           have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
-           have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
-           have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2
-             by sauto lq:on use:coqeq_sound_mutual.
-           exists (subst_PTm (scons (PProj PL a0) VarPTm) B2).
-           repeat split.
-           *** apply : Su_Transitive; eauto.
-               have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2
-                 by qauto use:regularity, E_Refl.
-               sfirstorder use:Su_Sig_Proj2.
-           *** apply : Su_Transitive; eauto.
-               sfirstorder use:Su_Sig_Proj2.
-           *** eauto using T_Proj2.
-           *** apply : T_Conv.
-               apply : T_Proj2; eauto.
-               move /E_Symmetric in haE.
-               move /regularity_sub0 in hSu21.
-               sfirstorder use:bind_inst.
-      * move => > /algo_metric_join; clear => ? ? ?. exfalso.
-        hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv.
-    (* ind ind case *)
-    + move => P a0 b0 c0.
-      case : b => //=;
-                   lazymatch goal with
-                   | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac
-                   | _ => move => > /algo_metric_join; clear => *; exfalso
-                   end.
-      * 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.
-      * 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.
-        move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0.
-        move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1.
-        have {}iha : a0 ∼ a1 by qauto l:on.
-        have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia.
-        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 Δ := 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.
-        have haE : Γ ⊢ a0 ≡ a1 ∈ PNat
-          by hauto l:on use:coqeq_sound_mutual.
-        have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual.
-        have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)).
-        eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero.
-        have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P
-          by eauto using T_Conv_E.
-        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 : (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.
-        eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto.
-        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. 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.
-        have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual.
-        exists (subst_PTm (scons a0 VarPTm) P).
-        repeat split => //=; eauto with wt.
-        apply : Su_Transitive.
-        apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt.
-        apply : Su_Eq. apply E_Refl. by apply T_Nat'.
-        apply : Su_Eq. apply hPE. by eauto.
-        move : hEInd. clear. hauto l:on use:regularity.
-Qed.
+(*     case : a h fb fa => //=. *)
+(*     + case : b => //=; move => > /algo_metric_join. *)
+(*       * 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. *)
+(*       * clear => ? ? _. exfalso. *)
+(*         hauto q:on use:REReds.var_inv, REReds.hne_ind_inv. *)
+(*     + case : b => //=; *)
+(*                    lazymatch goal with *)
+(*                    | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac *)
+(*                    | _ => move => > /algo_metric_join *)
+(*                    end. *)
+(*       * clear => *. exfalso. *)
+(*         hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv. *)
+(*       (* real case *) *)
+(*       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB. *)
+(*         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0. *)
+(*         move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1. *)
+(*         move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply]. *)
+(*         move => [j [hj [hal0 hal1]]]. *)
+(*         have /ih := hal0. *)
+(*         move /(_ hj). *)
+(*         move => [_ ihb]. *)
+(*         move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply]. *)
+(*         move => [hb01][C][hT0][hT1][hT2]hT3. *)
+(*         move /Sub_Bind_InvL : (hT0). *)
+(*         move => [i][A2][B2]hE. *)
+(*         have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by *)
+(*           eauto using Su_Eq, Su_Transitive. *)
+(*         have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by *)
+(*           eauto using Su_Eq, Su_Transitive. *)
+(*         have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
+(*         have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
+(*         have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
+(*         have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
+(*         move : ih (hj) hal1. repeat move/[apply]. *)
+(*         move => [ih _]. *)
+(*         move : ih (ha0') (ha1'); repeat move/[apply]. *)
+(*         move => iha. *)
+(*         split. qblast. *)
+(*         exists (subst_PTm (scons a0 VarPTm) B2). *)
+(*         split. *)
+(*         apply : Su_Transitive; eauto. *)
+(*         move /E_Refl in ha0. *)
+(*         hauto l:on use:Su_Pi_Proj2. *)
+(*         have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual. *)
+(*         split. *)
+(*         apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2). *)
+(*         move /regularity_sub0 : hSu10 => [i0]. *)
+(*         hauto l:on use:bind_inst. *)
+(*         hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl. *)
+(*         split. *)
+(*         by apply : T_App; eauto using T_Conv_E. *)
+(*         apply : T_Conv; eauto. *)
+(*         apply T_App with (A := A2) (B := B2); eauto. *)
+(*         apply : T_Conv_E; eauto. *)
+(*         move /E_Symmetric in h01. *)
+(*         move /regularity_sub0 : hSu20 => [i0]. *)
+(*         sfirstorder use:bind_inst. *)
+(*       * clear => ? ? ?. exfalso. *)
+(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv. *)
+(*       * clear => ? ? ?. exfalso. *)
+(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv. *)
+(*     + case : b => //=. *)
+(*       * move => i p h /algo_metric_join. clear => ? _ ?. exfalso. *)
+(*         hauto l:on use:REReds.hne_proj_inv, REReds.var_inv. *)
+(*       * move => > /algo_metric_join. clear => ? ? ?. exfalso. *)
+(*         hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv. *)
+(*       (* real case *) *)
+(*       * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0. *)
+(*         move : ha (hne0) (hne1); repeat move/[apply]. *)
+(*         move => [? [j []]]. subst. *)
+(*         move : ih; repeat move/[apply]. *)
+(*         move => [_ ih]. *)
+(*         case : p1. *)
+(*         ** move => Γ A B ha0 ha1. *)
+(*            move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
+(*            move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
+(*            move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply]. *)
+(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
+(*            split. sauto lq:on. *)
+(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
+(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive. *)
+(*            repeat split => //=. *)
+(*            hauto l:on use:Su_Sig_Proj1, Su_Transitive. *)
+(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
+(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
+(*         ** move => Γ A B  ha0  ha1. *)
+(*            move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
+(*            move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
+(*            move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply]. *)
+(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
+(*            split. sauto lq:on. *)
+(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
+(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1. *)
+(*            have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1. *)
+(*            have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
+(*            have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
+(*            have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2 *)
+(*              by sauto lq:on use:coqeq_sound_mutual. *)
+(*            exists (subst_PTm (scons (PProj PL a0) VarPTm) B2). *)
+(*            repeat split. *)
+(*            *** apply : Su_Transitive; eauto. *)
+(*                have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2 *)
+(*                  by qauto use:regularity, E_Refl. *)
+(*                sfirstorder use:Su_Sig_Proj2. *)
+(*            *** apply : Su_Transitive; eauto. *)
+(*                sfirstorder use:Su_Sig_Proj2. *)
+(*            *** eauto using T_Proj2. *)
+(*            *** apply : T_Conv. *)
+(*                apply : T_Proj2; eauto. *)
+(*                move /E_Symmetric in haE. *)
+(*                move /regularity_sub0 in hSu21. *)
+(*                sfirstorder use:bind_inst. *)
+(*       * move => > /algo_metric_join; clear => ? ? ?. exfalso. *)
+(*         hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv. *)
+(*     (* ind ind case *) *)
+(*     + move => P a0 b0 c0. *)
+(*       case : b => //=; *)
+(*                    lazymatch goal with *)
+(*                    | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac *)
+(*                    | _ => move => > /algo_metric_join; clear => *; exfalso *)
+(*                    end. *)
+(*       * 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. *)
+(*       * 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. *)
+(*         move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0. *)
+(*         move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1. *)
+(*         have {}iha : a0 ∼ a1 by qauto l:on. *)
+(*         have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia. *)
+(*         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 Δ := 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. *)
+(*         have haE : Γ ⊢ a0 ≡ a1 ∈ PNat *)
+(*           by hauto l:on use:coqeq_sound_mutual. *)
+(*         have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual. *)
+(*         have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)). *)
+(*         eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero. *)
+(*         have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P *)
+(*           by eauto using T_Conv_E. *)
+(*         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 : (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. *)
+(*         eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto. *)
+(*         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. 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. *)
+(*         have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual. *)
+(*         exists (subst_PTm (scons a0 VarPTm) P). *)
+(*         repeat split => //=; eauto with wt. *)
+(*         apply : Su_Transitive. *)
+(*         apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt. *)
+(*         apply : Su_Eq. apply E_Refl. by apply T_Nat'. *)
+(*         apply : Su_Eq. apply hPE. by eauto. *)
+(*         move : hEInd. clear. hauto l:on use:regularity. *)
+(* Qed. *)
 
-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_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 Γ (a b A : PTm) :
-  Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
-Proof.
-  move => h.
-  suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity.
-  eapply fundamental_theorem in h.
-  move /logrel.SemEq_SN_Join : h.
-  move => h.
-  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.
-  move => [va][vb][hva][hvb][nva][nvb]hj.
-  move /relations.rtc_nsteps : hva => [i hva].
-  move /relations.rtc_nsteps : hvb => [j hvb].
-  exists (i + j + size_PTm va + size_PTm vb).
-  hauto lq:on solve+:lia.
-Qed.
+(* Lemma coqeq_complete Γ (a b A : PTm) : *)
+(*   Γ ⊢ a ≡ b ∈ A -> a ⇔ b. *)
+(* Proof. *)
+(*   move => h. *)
+(*   suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity. *)
+(*   eapply fundamental_theorem in h. *)
+(*   move /logrel.SemEq_SN_Join : h. *)
+(*   move => h. *)
+(*   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. *)
+(*   move => [va][vb][hva][hvb][nva][nvb]hj. *)
+(*   move /relations.rtc_nsteps : hva => [i hva]. *)
+(*   move /relations.rtc_nsteps : hvb => [j hvb]. *)
+(*   exists (i + j + size_PTm va + size_PTm vb). *)
+(*   hauto lq:on solve+:lia. *)
+(* Qed. *)
 
 
 Reserved Notation "a ≪ b" (at level 70).
@@ -1804,295 +1798,295 @@ Scheme coqleq_ind := Induction for CoqLEq Sort Prop
 
 Combined Scheme coqleq_mutual from coqleq_ind, coqleq_r_ind.
 
-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.
+(* 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 k (a b : PTm) :
-  ishne a \/ ishne b ->
-  salgo_metric k a b ->
-  algo_metric k a b.
-Proof.
-  move => h.
-  move => [i][j][va][vb][hva][hvb][nva][nvb][hS]sz.
-  rewrite/ESub.R in hS.
-  move : hS => [va'][vb'][h0][h1]h2.
-  suff : va' = vb' by sauto lq:on.
-  have {}hva : rtc RERed.R a va by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds.
-  have {}hvb : rtc RERed.R b vb by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds.
-  apply REReds.FromEReds in h0, h1.
-  have : ishne va' \/ ishne vb' by
-    hauto lq:on rew:off use:@relations.rtc_transitive, REReds.hne_preservation.
-  hauto lq:on use:Sub1.hne_refl.
-Qed.
+(* Lemma salgo_metric_algo_metric k (a b : PTm) : *)
+(*   ishne a \/ ishne b -> *)
+(*   salgo_metric k a b -> *)
+(*   algo_metric k a b. *)
+(* Proof. *)
+(*   move => h. *)
+(*   move => [i][j][va][vb][hva][hvb][nva][nvb][hS]sz. *)
+(*   rewrite/ESub.R in hS. *)
+(*   move : hS => [va'][vb'][h0][h1]h2. *)
+(*   suff : va' = vb' by sauto lq:on. *)
+(*   have {}hva : rtc RERed.R a va by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds. *)
+(*   have {}hvb : rtc RERed.R b vb by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds. *)
+(*   apply REReds.FromEReds in h0, h1. *)
+(*   have : ishne va' \/ ishne vb' by *)
+(*     hauto lq:on rew:off use:@relations.rtc_transitive, REReds.hne_preservation. *)
+(*   hauto lq:on use:Sub1.hne_refl. *)
+(* Qed. *)
 
-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 => 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.
-    apply Su_Transitive with (B := PBind PPi A1 B0).
-    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 => 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.
-    apply Su_Transitive with (B := PBind PSig A0 B1).
-    apply : Su_Sig; eauto using E_Refl, Su_Eq.
-    apply : ihB; by eauto using ctx_eq_subst_one.
-    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 => 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 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 => 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. *)
+(*     apply Su_Transitive with (B := PBind PPi A1 B0). *)
+(*     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 => 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. *)
+(*     apply Su_Transitive with (B := PBind PSig A0 B1). *)
+(*     apply : Su_Sig; eauto using E_Refl, Su_Eq. *)
+(*     apply : ihB; by eauto using ctx_eq_subst_one. *)
+(*     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 => 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 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.
-  move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6.
-  case : a h0 => //=; try firstorder.
-  - inversion h0 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.AppAbs unfold:algo_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-    + sfirstorder qb:on use:ne_hne.
-  - inversion h0 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-  - inversion h0 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
-    + sfirstorder use:ne_hne.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
-    + sfirstorder use:ne_hne.
-    + sfirstorder use:ne_hne.
-Qed.
+(* 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. *)
+(*   move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6. *)
+(*   case : a h0 => //=; try firstorder. *)
+(*   - inversion h0 as [|A B C D E F]; subst. *)
+(*     hauto qb:on use:ne_hne. *)
+(*     inversion E; subst => /=. *)
+(*     + hauto lq:on use:HRed.AppAbs unfold:algo_metric solve+:lia. *)
+(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*     + sfirstorder qb:on use:ne_hne. *)
+(*   - inversion h0 as [|A B C D E F]; subst. *)
+(*     hauto qb:on use:ne_hne. *)
+(*     inversion E; subst => /=. *)
+(*     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia. *)
+(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*   - inversion h0 as [|A B C D E F]; subst. *)
+(*     hauto qb:on use:ne_hne. *)
+(*     inversion E; subst => /=. *)
+(*     + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia. *)
+(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*     + sfirstorder use:ne_hne. *)
+(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*     + sfirstorder use:ne_hne. *)
+(*     + sfirstorder use:ne_hne. *)
+(* Qed. *)
 
-Lemma CLE_HRedL (a a' b : PTm )  :
-  HRed.R a a' ->
-  a' ≪ b ->
-  a ≪ b.
-Proof.
-  hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
-Qed.
+(* Lemma CLE_HRedL (a a' b : PTm )  : *)
+(*   HRed.R a a' -> *)
+(*   a' ≪ b -> *)
+(*   a ≪ b. *)
+(* Proof. *)
+(*   hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R. *)
+(* Qed. *)
 
-Lemma CLE_HRedR (a a' b : PTm)  :
-  HRed.R a a' ->
-  b ≪ a' ->
-  b ≪ a.
-Proof.
-  hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
-Qed.
+(* Lemma CLE_HRedR (a a' b : PTm)  : *)
+(*   HRed.R a a' -> *)
+(*   b ≪ a' -> *)
+(*   b ≪ a. *)
+(* Proof. *)
+(*   hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R. *)
+(* Qed. *)
 
 
-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.
-  move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6.
-  case : b h1 => //=; try by firstorder.
-  - inversion 1 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto q:on use:HRed.AppAbs unfold:salgo_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:salgo_metric solve+:lia.
-    + sfirstorder qb:on use:ne_hne.
-  - inversion 1 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia.
-    + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia.
-  - inversion 1 as [|A B C D E F]; subst.
-    hauto qb:on use:ne_hne.
-    inversion E; subst => /=.
-    + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
-    + sfirstorder use:ne_hne.
-    + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia.
-    + sfirstorder use:ne_hne.
-    + sfirstorder use:ne_hne.
-Qed.
+(* 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. *)
+(*   move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6. *)
+(*   case : b h1 => //=; try by firstorder. *)
+(*   - inversion 1 as [|A B C D E F]; subst. *)
+(*     hauto qb:on use:ne_hne. *)
+(*     inversion E; subst => /=. *)
+(*     + hauto q:on use:HRed.AppAbs unfold:salgo_metric solve+:lia. *)
+(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:salgo_metric solve+:lia. *)
+(*     + sfirstorder qb:on use:ne_hne. *)
+(*   - inversion 1 as [|A B C D E F]; subst. *)
+(*     hauto qb:on use:ne_hne. *)
+(*     inversion E; subst => /=. *)
+(*     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia. *)
+(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*   - inversion 1 as [|A B C D E F]; subst. *)
+(*     hauto qb:on use:ne_hne. *)
+(*     inversion E; subst => /=. *)
+(*     + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia. *)
+(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*     + sfirstorder use:ne_hne. *)
+(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
+(*     + sfirstorder use:ne_hne. *)
+(*     + sfirstorder use:ne_hne. *)
+(* Qed. *)
 
-Lemma salgo_metric_sub k (a b : PTm) :
-  salgo_metric k a b ->
-  Sub.R a b.
-Proof.
-  rewrite /algo_metric.
-  move => [i][j][va][vb][h0][h1][h2][h3][[va' [vb' [hva [hvb hS]]]]]h5.
-  have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps.
-  have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps.
-  apply REReds.FromEReds in hva,hvb.
-  apply LoReds.ToRReds in h0,h1.
-  apply REReds.FromRReds in h0,h1.
-  rewrite /Sub.R. exists va', vb'. sfirstorder use:@relations.rtc_transitive.
-Qed.
+(* Lemma salgo_metric_sub k (a b : PTm) : *)
+(*   salgo_metric k a b -> *)
+(*   Sub.R a b. *)
+(* Proof. *)
+(*   rewrite /algo_metric. *)
+(*   move => [i][j][va][vb][h0][h1][h2][h3][[va' [vb' [hva [hvb hS]]]]]h5. *)
+(*   have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps. *)
+(*   have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps. *)
+(*   apply REReds.FromEReds in hva,hvb. *)
+(*   apply LoReds.ToRReds in h0,h1. *)
+(*   apply REReds.FromRReds in h0,h1. *)
+(*   rewrite /Sub.R. exists va', vb'. sfirstorder use:@relations.rtc_transitive. *)
+(* Qed. *)
 
-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.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
-  move : lored_nsteps_bind_inv h0 => /[apply].
-  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
-  move : lored_nsteps_bind_inv h1 => /[apply].
-  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
-  move /ESub.pi_inj : h4 => [? ?].
-  hauto qb:on solve+:lia.
-Qed.
+(* 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. *)
+(*   move => [i][j][va][vb][h0][h1][h2][h3][h4]h5. *)
+(*   move : lored_nsteps_bind_inv h0 => /[apply]. *)
+(*   move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst. *)
+(*   move : lored_nsteps_bind_inv h1 => /[apply]. *)
+(*   move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst. *)
+(*   move /ESub.pi_inj : h4 => [? ?]. *)
+(*   hauto qb:on solve+:lia. *)
+(* Qed. *)
 
-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.
-  move => [i][j][va][vb][h0][h1][h2][h3][h4]h5.
-  move : lored_nsteps_bind_inv h0 => /[apply].
-  move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst.
-  move : lored_nsteps_bind_inv h1 => /[apply].
-  move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst.
-  move /ESub.sig_inj : h4 => [? ?].
-  hauto qb:on solve+:lia.
-Qed.
+(* 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. *)
+(*   move => [i][j][va][vb][h0][h1][h2][h3][h4]h5. *)
+(*   move : lored_nsteps_bind_inv h0 => /[apply]. *)
+(*   move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst. *)
+(*   move : lored_nsteps_bind_inv h1 => /[apply]. *)
+(*   move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst. *)
+(*   move /ESub.sig_inj : h4 => [? ?]. *)
+(*   hauto qb:on solve+:lia. *)
+(* Qed. *)
 
-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 a b.
-  elim /Wf_nat.lt_wf_ind.
-  move => n ih.
-  move => a b /[dup] h /salgo_metric_case.
-  (* Cases where a and b can take steps *)
-  case; cycle 1.
-  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.
-  (* Cases where neither a nor b can take steps *)
-  case => fb; case => fa.
-  - case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-      * move => p0 A0 B0 _ p1 A1 B1 _.
-        move => h.
-        have ? : p1 = p0 by
-          hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj.
-        subst.
-        case : p0 h => //=.
-        ** move /salgo_metric_pi.
-           move => [j [hj [hA hB]]] Γ i.
-           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
-           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 : (cons A0 Γ) ⊢ B1 ∈ PUniv i
-             by hauto l:on.
-           eauto using ctx_eq_subst_one.
-        ** move /salgo_metric_sig.
-           move => [j [hj [hA hB]]] Γ i.
-           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
-           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 : (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.
-      * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf.
-      * move => _ > _ /salgo_metric_sub.
-        hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
-    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-      * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
-      * move => *. econstructor; eauto using rtc_refl.
-        hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong.
-      * hauto lq:on rew:off use:REReds.univ_inv, REReds.nat_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
-  (* Both cases are impossible *)
-    + case : b fb => //=.
-      * hauto lq:on use:T_AbsNat_Imp.
-      * hauto lq:on use:T_PairNat_Imp.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.bind_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto l:on.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
-    + move => ? ? Γ i /Zero_Inv.
-      clear.
-      move /synsub_to_usub => [_ [_ ]].
-      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
-    + move => ? _ _ Γ i /Suc_Inv => [[_]].
-      move /synsub_to_usub.
-      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
-  - have {}h : DJoin.R a b by
-       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join.
-    case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-    + hauto lq:on use:DJoin.hne_bind_noconf.
-    + hauto lq:on use:DJoin.hne_univ_noconf.
-    + hauto lq:on use:DJoin.hne_nat_noconf.
-    + move => _ _ Γ i _.
-      move /Zero_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-    + move => p _ _ Γ i _ /Suc_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-  - have {}h : DJoin.R b a by
-      hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric.
-    case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-    + hauto lq:on use:DJoin.hne_bind_noconf.
-    + hauto lq:on use:DJoin.hne_univ_noconf.
-    + hauto lq:on use:DJoin.hne_nat_noconf.
-    + move => _ _ Γ i.
-      move /Zero_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-    + move => p _ _ Γ i /Suc_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-  - move => Γ i ha hb.
-    econstructor; eauto using rtc_refl.
-    apply CLE_NeuNeu. move {ih}.
-    have {}h : algo_metric n a b by
-      hauto lq:on use:salgo_metric_algo_metric.
-    eapply coqeq_complete'; eauto.
-Qed.
+(* 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 a b. *)
+(*   elim /Wf_nat.lt_wf_ind. *)
+(*   move => n ih. *)
+(*   move => a b /[dup] h /salgo_metric_case. *)
+(*   (* Cases where a and b can take steps *) *)
+(*   case; cycle 1. *)
+(*   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. *)
+(*   (* Cases where neither a nor b can take steps *) *)
+(*   case => fb; case => fa. *)
+(*   - case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
+(*     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
+(*       * move => p0 A0 B0 _ p1 A1 B1 _. *)
+(*         move => h. *)
+(*         have ? : p1 = p0 by *)
+(*           hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj. *)
+(*         subst. *)
+(*         case : p0 h => //=. *)
+(*         ** move /salgo_metric_pi. *)
+(*            move => [j [hj [hA hB]]] Γ i. *)
+(*            move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0]. *)
+(*            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 : (cons A0 Γ) ⊢ B1 ∈ PUniv i *)
+(*              by hauto l:on. *)
+(*            eauto using ctx_eq_subst_one. *)
+(*         ** move /salgo_metric_sig. *)
+(*            move => [j [hj [hA hB]]] Γ i. *)
+(*            move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0]. *)
+(*            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 : (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. *)
+(*       * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf. *)
+(*       * move => _ > _ /salgo_metric_sub. *)
+(*         hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv inv:Sub1.R. *)
+(*       * hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R. *)
+(*     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
+(*       * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf. *)
+(*       * move => *. econstructor; eauto using rtc_refl. *)
+(*         hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong. *)
+(*       * hauto lq:on rew:off use:REReds.univ_inv, REReds.nat_inv, salgo_metric_sub inv:Sub1.R. *)
+(*       * hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R. *)
+(*       * hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R. *)
+(*   (* Both cases are impossible *) *)
+(*     + case : b fb => //=. *)
+(*       * hauto lq:on use:T_AbsNat_Imp. *)
+(*       * hauto lq:on use:T_PairNat_Imp. *)
+(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.bind_inv, salgo_metric_sub inv:Sub1.R. *)
+(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R. *)
+(*       * hauto l:on. *)
+(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R. *)
+(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R. *)
+(*     + move => ? ? Γ i /Zero_Inv. *)
+(*       clear. *)
+(*       move /synsub_to_usub => [_ [_ ]]. *)
+(*       hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R. *)
+(*     + move => ? _ _ Γ i /Suc_Inv => [[_]]. *)
+(*       move /synsub_to_usub. *)
+(*       hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R. *)
+(*   - have {}h : DJoin.R a b by *)
+(*        hauto lq:on use:salgo_metric_algo_metric, algo_metric_join. *)
+(*     case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
+(*     + hauto lq:on use:DJoin.hne_bind_noconf. *)
+(*     + hauto lq:on use:DJoin.hne_univ_noconf. *)
+(*     + hauto lq:on use:DJoin.hne_nat_noconf. *)
+(*     + move => _ _ Γ i _. *)
+(*       move /Zero_Inv. *)
+(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
+(*     + move => p _ _ Γ i _ /Suc_Inv. *)
+(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
+(*   - have {}h : DJoin.R b a by *)
+(*       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric. *)
+(*     case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
+(*     + hauto lq:on use:DJoin.hne_bind_noconf. *)
+(*     + hauto lq:on use:DJoin.hne_univ_noconf. *)
+(*     + hauto lq:on use:DJoin.hne_nat_noconf. *)
+(*     + move => _ _ Γ i. *)
+(*       move /Zero_Inv. *)
+(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
+(*     + move => p _ _ Γ i /Suc_Inv. *)
+(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
+(*   - move => Γ i ha hb. *)
+(*     econstructor; eauto using rtc_refl. *)
+(*     apply CLE_NeuNeu. move {ih}. *)
+(*     have {}h : algo_metric n a b by *)
+(*       hauto lq:on use:salgo_metric_algo_metric. *)
+(*     eapply coqeq_complete'; eauto. *)
+(* Qed. *)
 
-Lemma coqleq_complete Γ (a b : PTm) :
-  Γ ⊢ a ≲ b -> a ≪ b.
-Proof.
-  move => h.
-  suff : exists k, salgo_metric k a b by hauto lq:on use:coqleq_complete', regularity.
-  eapply fundamental_theorem in h.
-  move /logrel.SemLEq_SN_Sub : h.
-  move => h.
-  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.
-  move => [va][vb][hva][hvb][nva][nvb]hj.
-  move /relations.rtc_nsteps : hva => [i hva].
-  move /relations.rtc_nsteps : hvb => [j hvb].
-  exists (i + j + size_PTm va + size_PTm vb).
-  hauto lq:on solve+:lia.
-Qed.
+(* Lemma coqleq_complete Γ (a b : PTm) : *)
+(*   Γ ⊢ a ≲ b -> a ≪ b. *)
+(* Proof. *)
+(*   move => h. *)
+(*   suff : exists k, salgo_metric k a b by hauto lq:on use:coqleq_complete', regularity. *)
+(*   eapply fundamental_theorem in h. *)
+(*   move /logrel.SemLEq_SN_Sub : h. *)
+(*   move => h. *)
+(*   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. *)
+(*   move => [va][vb][hva][hvb][nva][nvb]hj. *)
+(*   move /relations.rtc_nsteps : hva => [i hva]. *)
+(*   move /relations.rtc_nsteps : hvb => [j hvb]. *)
+(*   exists (i + j + size_PTm va + size_PTm vb). *)
+(*   hauto lq:on solve+:lia. *)
+(* Qed. *)
 
-Lemma coqleq_sound : forall Γ (a b : PTm) i j,
-    Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv j -> a ≪ b -> Γ ⊢ a ≲ b.
-Proof.
-  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.
-  sfirstorder use:coqleq_sound_mutual.
-Qed.
+(* Lemma coqleq_sound : forall Γ (a b : PTm) i j, *)
+(*     Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv j -> a ≪ b -> Γ ⊢ a ≲ b. *)
+(* Proof. *)
+(*   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. *)
+(*   sfirstorder use:coqleq_sound_mutual. *)
+(* Qed. *)
diff --git a/theories/common.v b/theories/common.v
index 266d90a..18f4dfb 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -400,6 +400,8 @@ with salgo_dom_r : PTm -> PTm -> Prop :=
   salgo_dom_r a b.
 
 #[export]Hint Constructors salgo_dom salgo_dom_r : sdom.
+Scheme salgo_ind := Induction for salgo_dom Sort Prop
+  with salgor_ind := Induction for salgo_dom_r Sort Prop.
 
 Lemma hf_no_hred (a b : PTm) :
   ishf a ->
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 6ccee46..d51af9d 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -23,6 +23,14 @@ Proof.
   inversion 3; subst => //=.
 Qed.
 
+Lemma coqeq_neuneu' u0 u1 :
+  neuneu_nonconf u0 u1 ->
+  u0 ↔ u1 ->
+  u0 ∼ u1.
+Proof.
+  induction 2 => //=; destruct u => //=.
+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).
@@ -63,15 +71,15 @@ Proof.
   - sfirstorder.
   - move => i j /sumboolP ?.  subst.
     apply : CE_NeuNeu. apply CE_VarCong.
+  - 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 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].
@@ -117,14 +125,14 @@ Proof.
   - simp check_equal. done.
   - move => i j. move => h. have {h} : ~ nat_eqdec i j by done.
     case : nat_eqdec => //=. ce_solv.
+  - 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 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.
@@ -172,39 +180,42 @@ Qed.
 
 Ltac simp_sub := with_strategy opaque [check_equal] simpl in *.
 
+Combined Scheme salgo_dom_mutual from salgo_ind, salgor_ind.
+
 (* Reusing algo_dom results in an inefficient proof, but i'll brute force it so i can get the result more quickly *)
+Lemma check_sub_neuneu a b i a0 : check_sub a b (S_NeuNeu a b i a0) = check_equal a b a0.
+Proof. destruct a,b => //=. Qed.
+
+Lemma check_sub_conf a b n n0 i : check_sub a b (S_Conf a b n n0 i) = false.
+Proof. destruct a,b=>//=; ecrush inv:BTag. Qed.
+
+Hint Rewrite check_sub_neuneu check_sub_conf : ce_prop.
+
 Lemma check_sub_sound :
-  (forall a b (h : algo_dom a b), forall q, check_sub q a b h -> if q then a ⋖ b else b ⋖ a) /\
-    (forall a b (h : algo_dom_r a b), forall q, check_sub_r q a b h -> if q then a ≪ b else b ≪ a).
+  (forall a b (h : salgo_dom a b), check_sub a b h -> a ⋖ b) /\
+    (forall a b (h : salgo_dom_r a b), check_sub_r a b h -> a ≪ b).
 Proof.
-  apply algo_dom_mutual; try done.
-  - move => a [] //=; hauto qb:on.
-  - move => a0 a1 []//=; hauto qb:on.
+  apply salgo_dom_mutual; try done.
   - simpl. move => i j [];
     sauto lq:on use:Reflect.Nat_leb_le.
-  - move => p0 p1 A0 A1 B0 B1 a iha b ihb q.
-    case : p0; case : p1; try done; simp ce_prop.
-    sauto lqb:on.
-    sauto lqb:on.
-  - hauto l:on.
-  - move => i j q h.
-    have {}h : nat_eqdec i j by sfirstorder.
-    case : nat_eqdec h => //=; sauto lq:on.
-  - simp_sub.
-    sauto lqb:on use:coqeq_symmetric_mutual, coqeq_neuneu, check_equal_sound.
-  - simp_sub.
-    sauto lqb:on use:coqeq_symmetric_mutual, coqeq_neuneu, check_equal_sound.
-  - simp_sub.
-    sauto lqb:on use:coqeq_symmetric_mutual, coqeq_neuneu, check_equal_sound.
-  - move => a b n n0 i q h.
-    exfalso.
-    destruct a, b; try done; simp_sub; hauto lb:on use:check_equal_conf.
-  - move => a b doma ihdom q.
-    simp ce_prop. sauto lq:on.
-  - move => a a' b hr doma ihdom q.
-    simp ce_prop. move : ihdom => /[apply]. move {doma}.
-    sauto lq:on.
-  - move => a b b' n r dom ihdom q.
+  - move => A0 A1 B0 B1 s ihs s0 ihs0.
+    simp ce_prop.
+    hauto lqb:on ctrs:CoqLEq.
+  - move => A0 A1 B0 B1 s ihs s0 ihs0.
+    simp ce_prop.
+    hauto lqb:on ctrs:CoqLEq.
+  - qauto ctrs:CoqLEq.
+  - move => a b i a0.
+    simp ce_prop.
+    move => h. apply CLE_NeuNeu.
+    hauto lq:on use:check_equal_sound, coqeq_neuneu', coqeq_symmetric_mutual.
+  - move => a b n n0 i.
+    by simp ce_prop.
+  - move => a b h ih. simp ce_prop. hauto l:on.
+  - move => a a' b hr h ih.
+    simp ce_prop.
+    sauto lq:on rew:off.
+  - move => a b b' n r dom ihdom.
     simp ce_prop.
     move : ihdom => /[apply].
     move {dom}.
@@ -212,91 +223,51 @@ Proof.
 Qed.
 
 Lemma check_sub_complete :
-  (forall a b (h : algo_dom a b), forall q, check_sub q a b h = false -> if q then ~ a ⋖ b else ~ b ⋖ a) /\
-    (forall a b (h : algo_dom_r a b), forall q, check_sub_r q a b h = false -> if q then ~ a ≪ b else ~ b ≪ a).
+  (forall a b (h : salgo_dom a b), check_sub a b h = false -> ~ a ⋖ b) /\
+    (forall a b (h : salgo_dom_r a b), check_sub_r a b h = false -> ~ a ≪ b).
 Proof.
-  apply algo_dom_mutual; try first [done | hauto depth:4 lq:on inv:CoqLEq, CoqEq_Neu].
-  - move => i j q /=.
-    sauto lq:on rew:off use:PeanoNat.Nat.leb_le.
-  - move => p0 p1 A0 A1 B0 B1 a iha b ihb [].
-    + move => + h. inversion h; subst; simp ce_prop.
-      * move /nandP.
-        case.
-        by move => /negbTE {}/iha.
-        by move => /negbTE {}/ihb.
-      * move /nandP.
-        case.
-        by move => /negbTE {}/iha.
-        by move => /negbTE {}/ihb.
-      * inversion H.
-    + move => + h. inversion h; subst; simp ce_prop.
-      * move /nandP.
-        case.
-        by move => /negbTE {}/iha.
-        by move => /negbTE {}/ihb.
-      * move /nandP.
-        case.
-        by move => /negbTE {}/iha.
-        by move => /negbTE {}/ihb.
-      * inversion H.
-  - simp_sub.
-    sauto lq:on use:check_equal_complete.
-  - simp_sub.
-    move => p0 p1 u0 u1 i i0 a iha q.
+  apply salgo_dom_mutual; try first [done | hauto depth:4 lq:on inv:CoqLEq, CoqEq_Neu].
+  - move => i j /=.
+    move => + h. inversion h; subst => //=.
+    sfirstorder use:PeanoNat.Nat.leb_le.
+    hauto lq:on inv:CoqEq_Neu.
+  - move => A0 A1 B0 B1 s ihs s' ihs'. simp ce_prop.
     move /nandP.
     case.
-    move /nandP.
-    case => //.
-    by move /negP.
-    by move /negP.
-    move /negP.
-    move => h. eapply check_equal_complete in h.
-    sauto use:coqeq_symmetric_mutual inv:CoqLEq lq:on.
-  - move => u0 u1 a0 a1 i i0 a hdom ihdom hdom' ihdom'.
-    simp_sub.
+    move => /negbTE {}/ihs. hauto q:on inv:CoqLEq, CoqEq_Neu.
+    move => /negbTE {}/ihs'. hauto q:on inv:CoqLEq, CoqEq_Neu.
+  - move => A0 A1 B0 B1 s ihs s' ihs'. simp ce_prop.
     move /nandP.
     case.
-    move/nandP.
-    case.
-    by move/negP.
-    by move/negP.
-    move /negP.
-    move => h. eapply check_equal_complete in h.
-    sauto use:coqeq_symmetric_mutual inv:CoqLEq lq:on.
-  - move => P0 P1 u0 u1 b0 b1 c0 c1 i i0 dom ihdom dom' ihdom' dom'' ihdom'' dom''' ihdom''' q.
-    move /nandP.
-    case.
-    move /nandP.
-    case => //=.
-    by move/negP.
-    by move/negP.
-    move /negP => h. eapply check_equal_complete in h.
-    sauto use:coqeq_symmetric_mutual inv:CoqLEq lq:on.
-  - move => a b h ih q. simp ce_prop => {}/ih.
-    case : q => h0;
-    inversion 1; subst; hauto l:on use:algo_dom_no_hred, hreds_nf_refl.
-  - move => a a' b r dom ihdom q.
-    simp ce_prop => {}/ihdom.
-    case : q => h0.
-    inversion 1; subst.
-    inversion H0; subst. sfirstorder use:coqleq_no_hred.
-    have ? : a' = y by eauto using hred_deter. subst.
-    sauto lq:on.
-    inversion 1; subst.
-    inversion H1; subst. sfirstorder use:coqleq_no_hred.
-    have ? : a' = y by eauto using hred_deter. subst.
-    sauto lq:on.
-  - move => a b b' n r hr ih q.
-    simp ce_prop => {}/ih.
-    case : q.
-    + move => h. inversion 1; subst.
-      inversion H1; subst.
-      sfirstorder use:coqleq_no_hred.
-      have ? : b' = y by eauto using hred_deter. subst.
-      sauto lq:on.
-    + move => h. inversion 1; subst.
-      inversion H0; subst.
-      sfirstorder use:coqleq_no_hred.
-      have ? : b' = y by eauto using hred_deter. subst.
+    move => /negbTE {}/ihs. hauto q:on inv:CoqLEq, CoqEq_Neu.
+    move => /negbTE {}/ihs'. hauto q:on inv:CoqLEq, CoqEq_Neu.
+  - move => a b i hs. simp ce_prop.
+    move => + h. inversion h; subst => //=.
+    move => /negP h0.
+    eapply check_equal_complete in h0.
+    apply h0. by constructor.
+  - move => a b i0 i1 j. simp ce_prop.
+    move => _ h. inversion h; subst => //=.
+    hauto lq:on inv:CoqEq_Neu unfold:stm_conf.
+  - move => a b s ihs. simp ce_prop.
+    move => h0 h1. apply ihs =>//.
+    have [? ?] : HRed.nf a /\ HRed.nf b by hauto l:on use:salgo_dom_no_hred.
+    inversion h1; subst.
+    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:coqleq_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_sub_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:coqleq_no_hred.
+      sfirstorder unfold:HRed.nf.
+    + sauto lq:on use:hred_deter.
 Qed.

From 30caf750020c28a5b4e55d68c0354ddbdd5c22bb Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 19:58:19 -0400
Subject: [PATCH 193/210] Make progress on the refactored lemma

---
 theories/algorithmic.v | 1055 +++++++++++++++++++++-------------------
 1 file changed, 562 insertions(+), 493 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index de58fd4..ace67ac 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -252,22 +252,6 @@ Proof.
     apply : ctx_eq_subst_one; eauto.
 Qed.
 
-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 Γ p (A : PTm) B i :
-  Γ ⊢ PBind p A B ∈ PUniv i ->
-  Γ ⊢ A ∈ PUniv i /\ (cons A Γ) ⊢ B ∈ PUniv i.
-Proof.
-  move /Bind_Inv.
-  move => [i0][hA][hB]h.
-  move /synsub_to_usub : h => [_ [_ /Sub.univ_inj ? ]].
-  sfirstorder use:T_Univ_Raise.
-Qed.
-
 Lemma Abs_Pi_Inv Γ (a : PTm) A B :
   Γ ⊢ PAbs a ∈ PBind PPi A B ->
   (cons A Γ) ⊢ a ∈ B.
@@ -288,6 +272,48 @@ Proof.
   by asimpl; rewrite subst_scons_id.
 Qed.
 
+Lemma T_Abs_Neu_Inv Γ (a u : PTm) U :
+  Γ ⊢ PAbs a ∈ U ->
+  Γ ⊢ u ∈ U ->
+  exists Δ V, Δ ⊢ a ∈ V /\ Δ ⊢ PApp (ren_PTm shift u) (VarPTm var_zero)  ∈ V.
+Proof.
+  move => /[dup] ha' + hu.
+  move /Abs_Inv => [A0][B0][ha]hSu.
+  move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE.
+  have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
+  have ha'' : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
+  have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i
+    by hauto l:on use:regularity.
+  have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j
+      by qauto l:on use:Bind_Inv.
+  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.
+  apply T_Var; eauto using here.
+  rewrite -/ren_PTm. by asimpl; rewrite subst_scons_id.
+  exists Δ, B2. split => //.
+  by move /Abs_Pi_Inv in ha''.
+Qed.
+
+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 Γ p (A : PTm) B i :
+  Γ ⊢ PBind p A B ∈ PUniv i ->
+  Γ ⊢ A ∈ PUniv i /\ (cons A Γ) ⊢ B ∈ PUniv i.
+Proof.
+  move /Bind_Inv.
+  move => [i0][hA][hB]h.
+  move /synsub_to_usub : h => [_ [_ /Sub.univ_inj ? ]].
+  sfirstorder use:T_Univ_Raise.
+Qed.
+
 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.
@@ -422,6 +448,10 @@ Proof.
   hauto lq:on ctrs:rtc, CoqEq_R inv:CoqEq_R.
 Qed.
 
+Lemma CE_Nf a b :
+  a ↔ b -> a ⇔ b.
+Proof. hauto l:on ctrs:rtc. Qed.
+
 Scheme
   coqeq_neu_ind := Induction for CoqEq_Neu Sort Prop
   with coqeq_ind := Induction for CoqEq Sort Prop
@@ -1272,489 +1302,528 @@ Proof.
       hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
 Qed.
 
-(* Lemma coqeq_complete' k (a b : PTm ) : *)
-(*   (forall a b, algo_dom a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b) *)
+Lemma coqeq_complete' :
+  (forall a b, algo_dom a b -> DJoin.R 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)) /\
+    (forall a b, algo_dom_r a b -> DJoin.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b).
+  move => [:hhPairNeu hhAbsNeu].
+  apply algo_dom_mutual.
+  - move => a b h ih hj. split => //.
+    move => Γ A. move : T_Abs_Inv; repeat move/[apply].
+    move => [Δ][V][h0]h1.
+    have [? ?] : SN a /\ SN b by hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply CE_Nf. constructor. apply : ih; eauto using DJoin.abs_inj.
+  - abstract : hhAbsNeu.
+    move => a u hu ha iha hj => //.
+    split => //= Γ A.
+    move => + h. have ? : SN u  by hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    move : T_Abs_Neu_Inv h; repeat move/[apply].
+    move => [Δ][V][ha']hu'.
+    apply CE_Nf. constructor => //. apply : iha; eauto.
+    apply DJoin.abs_inj.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply : DJoin.transitive; eauto.
+    apply DJoin.symmetric. apply DJoin.FromEJoin. eexists. split. apply relations.rtc_once.
+    apply ERed.AppEta. apply rtc_refl.
+  - hauto q:on use:coqeq_symmetric_mutual, DJoin.symmetric, algo_dom_sym.
+  - move {hhAbsNeu hhPairNeu}.
+    admit.
+  - move {hhAbsNeu}. abstract : hhPairNeu.
+    admit.
+  - move {hhAbsNeu}.
+    move => a0 a1 u hu  ha iha hb ihb /DJoin.symmetric hj. split => // *.
+    eapply coqeq_symmetric_mutual.
+    eapply algo_dom_sym in ha, hb.
+    eapply hhPairNeu => //=; eauto; hauto lq:on use:DJoin.symmetric, coqeq_symmetric_mutual.
+  - move {hhPairNeu hhAbsNeu}. hauto l:on.
+  - move {hhPairNeu hhAbsNeu}.
+    move => a0 a1 ha iha /DJoin.suc_inj hj. split => //.
+    move => Γ A /Suc_Inv ? /Suc_Inv ?. apply CE_Nf. hauto lq:on ctrs:CoqEq.
+  - move => i j /DJoin.univ_inj ?. subst.
+    split => //. hauto l:on.
+  - move => p0 p1 A0 A1 B0 B1.
 
-(*   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 a b. *)
-(*   elim /Wf_nat.lt_wf_ind. *)
-(*   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 : 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 /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 : a b h. move => [:hnfneu]. *)
-(*   move => a b h. *)
-(*   case => fb; case => fa. *)
-(*   - split; last by sfirstorder use:hf_not_hne. *)
-(*     move {hnfneu}. *)
-(*     case : a h fb fa => //=. *)
-(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp. *)
-(*       move => a0 a1 ha0 _ _ Γ A wt0 wt1. *)
-(*       move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]]. *)
-(*       apply : CE_HRed; eauto using rtc_refl. *)
-(*       apply CE_AbsAbs. *)
-(*       suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto. *)
-(*       have ? : n > 0 by sauto solve+:lia. *)
-(*       exists (n - 1). split; first by lia. *)
-(*       move : (ha0). rewrite /algo_metric. *)
-(*       move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har. *)
-(*       apply lored_nsteps_abs_inv in hr0, hr1. *)
-(*       move : hr0 => [va' [hr00 hr01]]. *)
-(*       move : hr1 => [vb' [hr10 hr11]]. move {ih}. *)
-(*       exists i,j,va',vb'. subst. *)
-(*       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0. *)
-(*       repeat split =>//=. sfirstorder. *)
-(*       simpl in *; by lia. *)
-(*       move /algo_metric_join /DJoin.symmetric : ha0. *)
-(*       have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
-(*       move => /[dup] [[ha00 ha10]] []. *)
-(*       move : DJoin.abs_inj; repeat move/[apply]. *)
-(*       move : DJoin.standardization ha00 ha10; repeat move/[apply]. *)
-(*       (* this is awful *) *)
-(*       move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]]. *)
-(*       have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va' *)
-(*         by hauto lq:on use:@relations.rtc_nsteps. *)
-(*       have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb' *)
-(*         by hauto lq:on use:@relations.rtc_nsteps. *)
-(*       simpl in *. *)
-(*       have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst. *)
-(*       sfirstorder. *)
-(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp. *)
-(*       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1. *)
-(*       have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1) *)
-(*         by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
-(*       apply : CE_HRed; eauto using rtc_refl. *)
-(*       move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0. *)
-(*       move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1. *)
-(*       move /Sub_Bind_InvR :  (hSu0). *)
-(*       move => [i][A2][B2]hE. *)
-(*       have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2 *)
-(*         by eauto using Su_Transitive, Su_Eq. *)
-(*       have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 *)
-(*         by eauto using Su_Transitive, Su_Eq. *)
-(*       have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1. *)
-(*       have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1. *)
-(*       have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
-(*       have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
-(*       move /algo_metric_pair : h sn0 sn1; repeat move/[apply]. *)
-(*       move => [j][hj][ih0 ih1]. *)
-(*       have haE : a0 ⇔ a1 by hauto l:on. *)
-(*       apply CE_PairPair => //. *)
-(*       have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2 *)
-(*         by hauto l:on use:coqeq_sound_mutual. *)
-(*       have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2. *)
-(*       apply : T_Conv; eauto. *)
-(*       move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2. *)
-(*       eapply ih; cycle -1; eauto. *)
-(*       apply : T_Conv; eauto. *)
-(*       apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2). *)
-(*       move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2. *)
-(*       move : hE haE. clear. *)
-(*       move => h. *)
-(*       eapply regularity in h. *)
-(*       move : h => [_ [hB _]]. *)
-(*       eauto using bind_inst. *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on use:T_AbsBind_Imp. *)
-(*       * hauto lq:on rew:off use:T_PairBind_Imp. *)
-(*       * move => p1 A1 B1 p0 A0 B0. *)
-(*         move /algo_metric_bind. *)
-(*         move => [?][j [ih0 [ih1 ih2]]]_ _. subst. *)
-(*         move => Γ A hU0 hU1. *)
-(*         move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]]. *)
-(*         move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]]. *)
-(*         have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1) *)
-(*           by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia. *)
-(*         have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1) *)
-(*           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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1) *)
-(*           by hauto lq:on use:T_Univ_Raise solve+:lia. *)
-(*         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. *)
-(*         have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1) *)
-(*           by hauto l:on use:coqeq_sound_mutual. *)
-(*         apply : CE_HRed; eauto using rtc_refl. *)
-(*         constructor => //. *)
+  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 a b.
+  elim /Wf_nat.lt_wf_ind.
+  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 : 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 /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 : a b h. move => [:hnfneu].
+  move => a b h.
+  case => fb; case => fa.
+  - split; last by sfirstorder use:hf_not_hne.
+    move {hnfneu}.
+    case : a h fb fa => //=.
+    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp.
+      move => a0 a1 ha0 _ _ Γ A wt0 wt1.
+      move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]].
+      apply : CE_HRed; eauto using rtc_refl.
+      apply CE_AbsAbs.
+      suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto.
+      have ? : n > 0 by sauto solve+:lia.
+      exists (n - 1). split; first by lia.
+      move : (ha0). rewrite /algo_metric.
+      move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har.
+      apply lored_nsteps_abs_inv in hr0, hr1.
+      move : hr0 => [va' [hr00 hr01]].
+      move : hr1 => [vb' [hr10 hr11]]. move {ih}.
+      exists i,j,va',vb'. subst.
+      suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
+      repeat split =>//=. sfirstorder.
+      simpl in *; by lia.
+      move /algo_metric_join /DJoin.symmetric : ha0.
+      have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
+      move => /[dup] [[ha00 ha10]] [].
+      move : DJoin.abs_inj; repeat move/[apply].
+      move : DJoin.standardization ha00 ha10; repeat move/[apply].
+      (* this is awful *)
+      move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]].
+      have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va'
+        by hauto lq:on use:@relations.rtc_nsteps.
+      have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb'
+        by hauto lq:on use:@relations.rtc_nsteps.
+      simpl in *.
+      have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst.
+      sfirstorder.
+    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp.
+      move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
+      have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1)
+        by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
+      apply : CE_HRed; eauto using rtc_refl.
+      move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0.
+      move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1.
+      move /Sub_Bind_InvR :  (hSu0).
+      move => [i][A2][B2]hE.
+      have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2
+        by eauto using Su_Transitive, Su_Eq.
+      have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2
+        by eauto using Su_Transitive, Su_Eq.
+      have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1.
+      have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1.
+      have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
+      have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
+      move /algo_metric_pair : h sn0 sn1; repeat move/[apply].
+      move => [j][hj][ih0 ih1].
+      have haE : a0 ⇔ a1 by hauto l:on.
+      apply CE_PairPair => //.
+      have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2
+        by hauto l:on use:coqeq_sound_mutual.
+      have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2.
+      apply : T_Conv; eauto.
+      move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2.
+      eapply ih; cycle -1; eauto.
+      apply : T_Conv; eauto.
+      apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2).
+      move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2.
+      move : hE haE. clear.
+      move => h.
+      eapply regularity in h.
+      move : h => [_ [hB _]].
+      eauto using bind_inst.
+    + case : b => //=.
+      * hauto lq:on use:T_AbsBind_Imp.
+      * hauto lq:on rew:off use:T_PairBind_Imp.
+      * move => p1 A1 B1 p0 A0 B0.
+        move /algo_metric_bind.
+        move => [?][j [ih0 [ih1 ih2]]]_ _. subst.
+        move => Γ A hU0 hU1.
+        move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]].
+        move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]].
+        have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1)
+          by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia.
+        have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1)
+          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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1)
+          by hauto lq:on use:T_Univ_Raise solve+:lia.
+        have {}hB1 : (cons A1 Γ) ⊢ B1 ∈ PUniv (max i0 i1)
+          by hauto lq:on use:T_Univ_Raise solve+:lia.
 
-(*         eapply ih; eauto. *)
-(*         apply : ctx_eq_subst_one; eauto. *)
-(*         eauto using Su_Eq. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on use:DJoin.bind_univ_noconf. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin. *)
-(*       * move => > /algo_metric_join. *)
-(*         clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv. *)
-(*       * move => > /algo_metric_join. clear. *)
-(*         hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv. *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on use:T_AbsUniv_Imp. *)
-(*       * hauto lq:on use:T_PairUniv_Imp. *)
-(*       * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric. *)
-(*       * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj. *)
-(*         subst. *)
-(*         hauto l:on. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin. *)
-(*       * move => > /algo_metric_join. *)
-(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv. *)
-(*     + case : b => //=. *)
-(*       * qauto l:on use:T_AbsNat_Imp. *)
-(*       * qauto l:on use:T_PairNat_Imp. *)
-(*       * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf. *)
-(*       * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf. *)
-(*       * hauto l:on. *)
-(*       * move /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv. *)
-(*     (* Zero *) *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on rew:off use:T_AbsZero_Imp. *)
-(*       * hauto lq: on use: T_PairZero_Imp. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv. *)
-(*       * hauto l:on. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv. *)
-(*     (* Suc *) *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on rew:off use:T_AbsSuc_Imp. *)
-(*       * hauto lq:on use:T_PairSuc_Imp. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv. *)
-(*       * move => a0 a1 h _ _. *)
-(*         move /algo_metric_suc : h => [j [h4 h5]]. *)
-(*         move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3]. *)
-(*         move : ih h4 h5;do!move/[apply]. *)
-(*         move => [ih _]. *)
-(*         move : ih h0 h2;do!move/[apply]. *)
-(*         move => h. apply : CE_HRed; eauto using rtc_refl. *)
-(*         by constructor. *)
-(*   - move : a b h fb fa. abstract : hnfneu. *)
-(*     move => + b. *)
-(*     case : b => //=. *)
-(*     (* NeuAbs *) *)
-(*     + move => a u halg _ nea. split => // Γ A hu /[dup] ha. *)
-(*       move /Abs_Inv => [A0][B0][_]hSu. *)
-(*       move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE. *)
-(*       have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
-(*       have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
-(*       have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i *)
-(*         by hauto l:on use:regularity. *)
-(*       have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j *)
-(*           by qauto l:on use:Bind_Inv. *)
-(*       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. *)
-(*       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. *)
-(*       apply : CE_HRed; eauto using rtc_refl. *)
-(*       eapply CE_NeuAbs => //=. *)
-(*       eapply ih; eauto. *)
-(*     (* NeuPair *) *)
-(*     + move => a0 b0 u halg _ neu. *)
-(*       split => // Γ A hu /[dup] wt. *)
-(*       move /Pair_Inv => [A0][B0][ha0][hb0]hU. *)
-(*       move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE. *)
-(*       have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on. *)
-(*       have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
-(*       move /Pair_Sig_Inv : wt => [{}ha0 {}hb0]. *)
-(*       have /T_Proj1 huL := hu. *)
-(*       have /T_Proj2 {hu} huR := hu. *)
-(*       move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]]. *)
-(*       have heL : PProj PL u ⇔ a0 by hauto l:on. *)
-(*       eapply CE_HRed; eauto using rtc_refl. *)
-(*       apply CE_NeuPair; eauto. *)
-(*       eapply ih; eauto. *)
-(*       apply : T_Conv; eauto. *)
-(*       have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i *)
-(*         by hauto l:on use:regularity. *)
-(*       have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by *)
-(*         hauto l:on use:coqeq_sound_mutual. *)
-(*       hauto l:on use:bind_inst. *)
-(*     (* NeuBind: Impossible *) *)
-(*     + move => b p p0 a /algo_metric_join h _ h0. exfalso. *)
-(*       move : h h0. clear. *)
-(*       move /Sub.FromJoin. *)
-(*       hauto l:on use:Sub.hne_bind_noconf. *)
-(*     (* NeuUniv: Impossible *) *)
-(*     + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join. *)
-(*     + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join. *)
-(*     (* Zero *) *)
-(*     + case => //=. *)
-(*       * move => i /algo_metric_join. clear. *)
-(*         hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv. *)
-(*       * move => >/algo_metric_join. clear. *)
-(*         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. *)
-(*       * move => >/algo_metric_join. clear. *)
-(*         move => h _ h2. exfalso. *)
-(*         hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv. *)
-(*     (* Suc *) *)
-(*     + move => a0. *)
-(*       case => //=; move => >/algo_metric_join; clear. *)
-(*       * 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. *)
-(*       * 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 Γ 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. *)
+        have ihA : A0 ⇔ A1 by hauto l:on.
+        have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1)
+          by hauto l:on use:coqeq_sound_mutual.
+        apply : CE_HRed; eauto using rtc_refl.
+        constructor => //.
 
-(*     case : a h fb fa => //=. *)
-(*     + case : b => //=; move => > /algo_metric_join. *)
-(*       * 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. *)
-(*       * clear => ? ? _. exfalso. *)
-(*         hauto q:on use:REReds.var_inv, REReds.hne_ind_inv. *)
-(*     + case : b => //=; *)
-(*                    lazymatch goal with *)
-(*                    | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac *)
-(*                    | _ => move => > /algo_metric_join *)
-(*                    end. *)
-(*       * clear => *. exfalso. *)
-(*         hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv. *)
-(*       (* real case *) *)
-(*       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB. *)
-(*         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0. *)
-(*         move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1. *)
-(*         move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply]. *)
-(*         move => [j [hj [hal0 hal1]]]. *)
-(*         have /ih := hal0. *)
-(*         move /(_ hj). *)
-(*         move => [_ ihb]. *)
-(*         move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply]. *)
-(*         move => [hb01][C][hT0][hT1][hT2]hT3. *)
-(*         move /Sub_Bind_InvL : (hT0). *)
-(*         move => [i][A2][B2]hE. *)
-(*         have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by *)
-(*           eauto using Su_Eq, Su_Transitive. *)
-(*         have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by *)
-(*           eauto using Su_Eq, Su_Transitive. *)
-(*         have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
-(*         have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
-(*         have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
-(*         have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
-(*         move : ih (hj) hal1. repeat move/[apply]. *)
-(*         move => [ih _]. *)
-(*         move : ih (ha0') (ha1'); repeat move/[apply]. *)
-(*         move => iha. *)
-(*         split. qblast. *)
-(*         exists (subst_PTm (scons a0 VarPTm) B2). *)
-(*         split. *)
-(*         apply : Su_Transitive; eauto. *)
-(*         move /E_Refl in ha0. *)
-(*         hauto l:on use:Su_Pi_Proj2. *)
-(*         have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual. *)
-(*         split. *)
-(*         apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2). *)
-(*         move /regularity_sub0 : hSu10 => [i0]. *)
-(*         hauto l:on use:bind_inst. *)
-(*         hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl. *)
-(*         split. *)
-(*         by apply : T_App; eauto using T_Conv_E. *)
-(*         apply : T_Conv; eauto. *)
-(*         apply T_App with (A := A2) (B := B2); eauto. *)
-(*         apply : T_Conv_E; eauto. *)
-(*         move /E_Symmetric in h01. *)
-(*         move /regularity_sub0 : hSu20 => [i0]. *)
-(*         sfirstorder use:bind_inst. *)
-(*       * clear => ? ? ?. exfalso. *)
-(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv. *)
-(*       * clear => ? ? ?. exfalso. *)
-(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv. *)
-(*     + case : b => //=. *)
-(*       * move => i p h /algo_metric_join. clear => ? _ ?. exfalso. *)
-(*         hauto l:on use:REReds.hne_proj_inv, REReds.var_inv. *)
-(*       * move => > /algo_metric_join. clear => ? ? ?. exfalso. *)
-(*         hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv. *)
-(*       (* real case *) *)
-(*       * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0. *)
-(*         move : ha (hne0) (hne1); repeat move/[apply]. *)
-(*         move => [? [j []]]. subst. *)
-(*         move : ih; repeat move/[apply]. *)
-(*         move => [_ ih]. *)
-(*         case : p1. *)
-(*         ** move => Γ A B ha0 ha1. *)
-(*            move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
-(*            move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
-(*            move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply]. *)
-(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
-(*            split. sauto lq:on. *)
-(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
-(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive. *)
-(*            repeat split => //=. *)
-(*            hauto l:on use:Su_Sig_Proj1, Su_Transitive. *)
-(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
-(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
-(*         ** move => Γ A B  ha0  ha1. *)
-(*            move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
-(*            move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
-(*            move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply]. *)
-(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
-(*            split. sauto lq:on. *)
-(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
-(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1. *)
-(*            have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1. *)
-(*            have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
-(*            have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
-(*            have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2 *)
-(*              by sauto lq:on use:coqeq_sound_mutual. *)
-(*            exists (subst_PTm (scons (PProj PL a0) VarPTm) B2). *)
-(*            repeat split. *)
-(*            *** apply : Su_Transitive; eauto. *)
-(*                have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2 *)
-(*                  by qauto use:regularity, E_Refl. *)
-(*                sfirstorder use:Su_Sig_Proj2. *)
-(*            *** apply : Su_Transitive; eauto. *)
-(*                sfirstorder use:Su_Sig_Proj2. *)
-(*            *** eauto using T_Proj2. *)
-(*            *** apply : T_Conv. *)
-(*                apply : T_Proj2; eauto. *)
-(*                move /E_Symmetric in haE. *)
-(*                move /regularity_sub0 in hSu21. *)
-(*                sfirstorder use:bind_inst. *)
-(*       * move => > /algo_metric_join; clear => ? ? ?. exfalso. *)
-(*         hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv. *)
-(*     (* ind ind case *) *)
-(*     + move => P a0 b0 c0. *)
-(*       case : b => //=; *)
-(*                    lazymatch goal with *)
-(*                    | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac *)
-(*                    | _ => move => > /algo_metric_join; clear => *; exfalso *)
-(*                    end. *)
-(*       * 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. *)
-(*       * 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. *)
-(*         move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0. *)
-(*         move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1. *)
-(*         have {}iha : a0 ∼ a1 by qauto l:on. *)
-(*         have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia. *)
-(*         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 Δ := 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. *)
-(*         have haE : Γ ⊢ a0 ≡ a1 ∈ PNat *)
-(*           by hauto l:on use:coqeq_sound_mutual. *)
-(*         have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual. *)
-(*         have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)). *)
-(*         eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero. *)
-(*         have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P *)
-(*           by eauto using T_Conv_E. *)
-(*         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 : (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. *)
-(*         eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto. *)
-(*         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. 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. *)
-(*         have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual. *)
-(*         exists (subst_PTm (scons a0 VarPTm) P). *)
-(*         repeat split => //=; eauto with wt. *)
-(*         apply : Su_Transitive. *)
-(*         apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt. *)
-(*         apply : Su_Eq. apply E_Refl. by apply T_Nat'. *)
-(*         apply : Su_Eq. apply hPE. by eauto. *)
-(*         move : hEInd. clear. hauto l:on use:regularity. *)
-(* Qed. *)
+        eapply ih; eauto.
+        apply : ctx_eq_subst_one; eauto.
+        eauto using Su_Eq.
+      * move => > /algo_metric_join.
+        hauto lq:on use:DJoin.bind_univ_noconf.
+      * move => > /algo_metric_join.
+        hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin.
+      * move => > /algo_metric_join.
+        clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv.
+      * move => > /algo_metric_join. clear.
+        hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv.
+    + case : b => //=.
+      * hauto lq:on use:T_AbsUniv_Imp.
+      * hauto lq:on use:T_PairUniv_Imp.
+      * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric.
+      * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj.
+        subst.
+        hauto l:on.
+      * move => > /algo_metric_join.
+        hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin.
+      * move => > /algo_metric_join.
+        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv.
+      * move => > /algo_metric_join.
+        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv.
+    + case : b => //=.
+      * qauto l:on use:T_AbsNat_Imp.
+      * qauto l:on use:T_PairNat_Imp.
+      * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf.
+      * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf.
+      * hauto l:on.
+      * move /algo_metric_join.
+        hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv.
+    (* Zero *)
+    + case : b => //=.
+      * hauto lq:on rew:off use:T_AbsZero_Imp.
+      * hauto lq: on use: T_PairZero_Imp.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv.
+      * hauto l:on.
+      * move =>> /algo_metric_join.
+        hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv.
+    (* Suc *)
+    + case : b => //=.
+      * hauto lq:on rew:off use:T_AbsSuc_Imp.
+      * hauto lq:on use:T_PairSuc_Imp.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv.
+      * move => > /algo_metric_join.
+        hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv.
+      * move => a0 a1 h _ _.
+        move /algo_metric_suc : h => [j [h4 h5]].
+        move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3].
+        move : ih h4 h5;do!move/[apply].
+        move => [ih _].
+        move : ih h0 h2;do!move/[apply].
+        move => h. apply : CE_HRed; eauto using rtc_refl.
+        by constructor.
+  - move : a b h fb fa. abstract : hnfneu.
+    move => + b.
+    case : b => //=.
+    (* NeuAbs *)
+    + move => a u halg _ nea. split => // Γ A hu /[dup] ha.
+      move /Abs_Inv => [A0][B0][_]hSu.
+      move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE.
+      have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
+      have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
+      have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i
+        by hauto l:on use:regularity.
+      have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j
+          by qauto l:on use:Bind_Inv.
+      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.
+      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.
+      apply : CE_HRed; eauto using rtc_refl.
+      eapply CE_NeuAbs => //=.
+      eapply ih; eauto.
+    (* NeuPair *)
+    + move => a0 b0 u halg _ neu.
+      split => // Γ A hu /[dup] wt.
+      move /Pair_Inv => [A0][B0][ha0][hb0]hU.
+      move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE.
+      have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on.
+      have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+      move /Pair_Sig_Inv : wt => [{}ha0 {}hb0].
+      have /T_Proj1 huL := hu.
+      have /T_Proj2 {hu} huR := hu.
+      move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]].
+      have heL : PProj PL u ⇔ a0 by hauto l:on.
+      eapply CE_HRed; eauto using rtc_refl.
+      apply CE_NeuPair; eauto.
+      eapply ih; eauto.
+      apply : T_Conv; eauto.
+      have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i
+        by hauto l:on use:regularity.
+      have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by
+        hauto l:on use:coqeq_sound_mutual.
+      hauto l:on use:bind_inst.
+    (* NeuBind: Impossible *)
+    + move => b p p0 a /algo_metric_join h _ h0. exfalso.
+      move : h h0. clear.
+      move /Sub.FromJoin.
+      hauto l:on use:Sub.hne_bind_noconf.
+    (* NeuUniv: Impossible *)
+    + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join.
+    + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join.
+    (* Zero *)
+    + case => //=.
+      * move => i /algo_metric_join. clear.
+        hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv.
+      * move => >/algo_metric_join. clear.
+        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.
+      * move => >/algo_metric_join. clear.
+        move => h _ h2. exfalso.
+        hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv.
+    (* Suc *)
+    + move => a0.
+      case => //=; move => >/algo_metric_join; clear.
+      * 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.
+      * 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 Γ 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.
 
-(* Lemma coqeq_sound : forall Γ (a b : PTm) A, *)
-(*     Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b -> Γ ⊢ a ≡ b ∈ A. *)
-(* Proof. sfirstorder use:coqeq_sound_mutual. Qed. *)
+    case : a h fb fa => //=.
+    + case : b => //=; move => > /algo_metric_join.
+      * 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.
+      * clear => ? ? _. exfalso.
+        hauto q:on use:REReds.var_inv, REReds.hne_ind_inv.
+    + case : b => //=;
+                   lazymatch goal with
+                   | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac
+                   | _ => move => > /algo_metric_join
+                   end.
+      * clear => *. exfalso.
+        hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv.
+      (* real case *)
+      * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
+        move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
+        move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1.
+        move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply].
+        move => [j [hj [hal0 hal1]]].
+        have /ih := hal0.
+        move /(_ hj).
+        move => [_ ihb].
+        move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply].
+        move => [hb01][C][hT0][hT1][hT2]hT3.
+        move /Sub_Bind_InvL : (hT0).
+        move => [i][A2][B2]hE.
+        have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by
+          eauto using Su_Eq, Su_Transitive.
+        have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by
+          eauto using Su_Eq, Su_Transitive.
+        have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
+        have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
+        have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
+        have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
+        move : ih (hj) hal1. repeat move/[apply].
+        move => [ih _].
+        move : ih (ha0') (ha1'); repeat move/[apply].
+        move => iha.
+        split. qblast.
+        exists (subst_PTm (scons a0 VarPTm) B2).
+        split.
+        apply : Su_Transitive; eauto.
+        move /E_Refl in ha0.
+        hauto l:on use:Su_Pi_Proj2.
+        have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
+        split.
+        apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
+        move /regularity_sub0 : hSu10 => [i0].
+        hauto l:on use:bind_inst.
+        hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
+        split.
+        by apply : T_App; eauto using T_Conv_E.
+        apply : T_Conv; eauto.
+        apply T_App with (A := A2) (B := B2); eauto.
+        apply : T_Conv_E; eauto.
+        move /E_Symmetric in h01.
+        move /regularity_sub0 : hSu20 => [i0].
+        sfirstorder use:bind_inst.
+      * clear => ? ? ?. exfalso.
+        hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
+      * clear => ? ? ?. exfalso.
+        hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv.
+    + case : b => //=.
+      * move => i p h /algo_metric_join. clear => ? _ ?. exfalso.
+        hauto l:on use:REReds.hne_proj_inv, REReds.var_inv.
+      * move => > /algo_metric_join. clear => ? ? ?. exfalso.
+        hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv.
+      (* real case *)
+      * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0.
+        move : ha (hne0) (hne1); repeat move/[apply].
+        move => [? [j []]]. subst.
+        move : ih; repeat move/[apply].
+        move => [_ ih].
+        case : p1.
+        ** move => Γ A B ha0 ha1.
+           move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0.
+           move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1.
+           move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply].
+           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
+           split. sauto lq:on.
+           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
+           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
+             by eauto using Su_Transitive, Su_Eq.
+           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
+             by eauto using Su_Transitive, Su_Eq.
+           exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive.
+           repeat split => //=.
+           hauto l:on use:Su_Sig_Proj1, Su_Transitive.
+           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
+           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
+        ** move => Γ A B  ha0  ha1.
+           move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0.
+           move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1.
+           move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply].
+           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
+           split. sauto lq:on.
+           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
+           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
+             by eauto using Su_Transitive, Su_Eq.
+           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
+             by eauto using Su_Transitive, Su_Eq.
+           have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1.
+           have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1.
+           have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+           have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+           have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2
+             by sauto lq:on use:coqeq_sound_mutual.
+           exists (subst_PTm (scons (PProj PL a0) VarPTm) B2).
+           repeat split.
+           *** apply : Su_Transitive; eauto.
+               have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2
+                 by qauto use:regularity, E_Refl.
+               sfirstorder use:Su_Sig_Proj2.
+           *** apply : Su_Transitive; eauto.
+               sfirstorder use:Su_Sig_Proj2.
+           *** eauto using T_Proj2.
+           *** apply : T_Conv.
+               apply : T_Proj2; eauto.
+               move /E_Symmetric in haE.
+               move /regularity_sub0 in hSu21.
+               sfirstorder use:bind_inst.
+      * move => > /algo_metric_join; clear => ? ? ?. exfalso.
+        hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv.
+    (* ind ind case *)
+    + move => P a0 b0 c0.
+      case : b => //=;
+                   lazymatch goal with
+                   | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac
+                   | _ => move => > /algo_metric_join; clear => *; exfalso
+                   end.
+      * 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.
+      * 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.
+        move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0.
+        move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1.
+        have {}iha : a0 ∼ a1 by qauto l:on.
+        have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia.
+        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 Δ := 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.
+        have haE : Γ ⊢ a0 ≡ a1 ∈ PNat
+          by hauto l:on use:coqeq_sound_mutual.
+        have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual.
+        have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)).
+        eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero.
+        have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P
+          by eauto using T_Conv_E.
+        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 : (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.
+        eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto.
+        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. 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.
+        have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual.
+        exists (subst_PTm (scons a0 VarPTm) P).
+        repeat split => //=; eauto with wt.
+        apply : Su_Transitive.
+        apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt.
+        apply : Su_Eq. apply E_Refl. by apply T_Nat'.
+        apply : Su_Eq. apply hPE. by eauto.
+        move : hEInd. clear. hauto l:on use:regularity.
+Qed.
 
-(* Lemma coqeq_complete Γ (a b A : PTm) : *)
-(*   Γ ⊢ a ≡ b ∈ A -> a ⇔ b. *)
-(* Proof. *)
-(*   move => h. *)
-(*   suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity. *)
-(*   eapply fundamental_theorem in h. *)
-(*   move /logrel.SemEq_SN_Join : h. *)
-(*   move => h. *)
-(*   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. *)
-(*   move => [va][vb][hva][hvb][nva][nvb]hj. *)
-(*   move /relations.rtc_nsteps : hva => [i hva]. *)
-(*   move /relations.rtc_nsteps : hvb => [j hvb]. *)
-(*   exists (i + j + size_PTm va + size_PTm vb). *)
-(*   hauto lq:on solve+:lia. *)
-(* Qed. *)
+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 Γ (a b A : PTm) :
+  Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
+Proof.
+  move => h.
+  suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity.
+  eapply fundamental_theorem in h.
+  move /logrel.SemEq_SN_Join : h.
+  move => h.
+  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.
+  move => [va][vb][hva][hvb][nva][nvb]hj.
+  move /relations.rtc_nsteps : hva => [i hva].
+  move /relations.rtc_nsteps : hvb => [j hvb].
+  exists (i + j + size_PTm va + size_PTm vb).
+  hauto lq:on solve+:lia.
+Qed.
 
 
 Reserved Notation "a ≪ b" (at level 70).

From d1771adc48b3ec0da063990e7db22834194f5fc3 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 22:44:42 -0400
Subject: [PATCH 194/210] Use hne and hf instead HRed.nf

---
 theories/algorithmic.v        | 1122 +++++++++++++++++++--------------
 theories/common.v             |   10 +-
 theories/executable.v         |    2 +-
 theories/executable_correct.v |    3 -
 4 files changed, 655 insertions(+), 482 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index ace67ac..eb93263 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1302,6 +1302,17 @@ Proof.
       hauto lq:on use:term_metric_ind, algo_dom_r_algo_dom db:adom.
 Qed.
 
+Lemma ce_neu_neu_helper a b :
+  ishne a -> ishne b ->
+  (forall Γ A B, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ B -> a ∼ b /\ exists C, Γ ⊢ C ≲ A /\ Γ ⊢ C ≲ B /\ Γ ⊢ a ∈ C /\ Γ ⊢ b ∈ C) -> (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. sauto lq:on. Qed.
+
+Lemma hne_ind_inj P0 P1 u0 u1 b0 b1 c0 c1 :
+  ishne u0 -> ishne u1 ->
+  DJoin.R (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1) ->
+  DJoin.R P0 P1 /\ DJoin.R u0 u1 /\ DJoin.R b0 b1 /\ DJoin.R c0 c1.
+Proof. hauto q:on use:REReds.hne_ind_inv. Qed.
+
 Lemma coqeq_complete' :
   (forall a b, algo_dom a b -> DJoin.R 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)) /\
     (forall a b, algo_dom_r a b -> DJoin.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b).
@@ -1341,489 +1352,654 @@ Lemma coqeq_complete' :
     move => Γ A /Suc_Inv ? /Suc_Inv ?. apply CE_Nf. hauto lq:on ctrs:CoqEq.
   - move => i j /DJoin.univ_inj ?. subst.
     split => //. hauto l:on.
-  - move => p0 p1 A0 A1 B0 B1.
+  - move => {hhPairNeu hhAbsNeu} p0 p1 A0 A1 B0 B1.
+    move => hA ihA hB ihB /DJoin.bind_inj. move => [?][hjA]hjB. subst.
+    split => // Γ A.
+    move => hbd0 hbd1.
+    have {hbd0} : exists i, Γ ⊢ PBind p1 A0 B0 ∈ PUniv i by move /Bind_Inv in hbd0; qauto use:T_Bind.
+    move => [i] => hbd0.
+    have {hbd1} : exists i, Γ ⊢ PBind p1 A1 B1 ∈ PUniv i by move /Bind_Inv in hbd1; qauto use:T_Bind.
+    move => [j] => hbd1.
+    have /Bind_Univ_Inv {hbd0} [? ?] : Γ ⊢ PBind p1 A0 B0 ∈ PUniv (max i j) by hauto lq:on use:T_Univ_Raise solve+:lia.
+    have /Bind_Univ_Inv {hbd1} [? ?] : Γ ⊢ PBind p1 A1 B1 ∈ PUniv (max i j) by hauto lq:on use:T_Univ_Raise solve+:lia.
+    move => [:eqa].
+    apply CE_Nf. constructor; first by abstract : eqa; eauto.
+    eapply ihB; eauto. apply : ctx_eq_subst_one; eauto.
+    apply : Su_Eq; eauto. sfirstorder use:coqeq_sound_mutual.
+  - hauto l:on.
+  - move => {hhAbsNeu hhPairNeu} i j /DJoin.var_inj ?. subst. apply ce_neu_neu_helper => // Γ A B.
+    move /Var_Inv => [h [A0 [h0 h1]]].
+    move /Var_Inv => [h' [A1 [h0' h1']]].
+    split. by constructor.
+    suff : A0 = A1 by hauto lq:on db:wt.
+    eauto using lookup_deter.
+  - move => u0 u1 a0 a1 neu0 neu1 domu ihu doma iha. move /DJoin.hne_app_inj /(_ neu0 neu1) => [hju hja].
+    apply ce_neu_neu_helper => //= Γ A B.
+    move /App_Inv => [A0][B0][hb0][ha0]hS0.
+    move /App_Inv => [A1][B1][hb1][ha1]hS1.
+    move /(_ hju) : ihu.
+    move => [_ ihb].
+    move : ihb (neu0) (neu1) hb0 hb1. repeat move/[apply].
+    move => [hb01][C][hT0][hT1][hT2]hT3.
+    move /Sub_Bind_InvL : (hT0).
+    move => [i][A2][B2]hE.
+    have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by
+      eauto using Su_Eq, Su_Transitive.
+    have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by
+      eauto using Su_Eq, Su_Transitive.
+    have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
+    have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
+    have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
+    have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
+    move : iha hja. repeat move/[apply].
+    move => iha.
+    move : iha (ha0') (ha1'); repeat move/[apply].
+    move => iha.
+    split. sauto lq:on.
+    exists (subst_PTm (scons a0 VarPTm) B2).
+    split.
+    apply : Su_Transitive; eauto.
+    move /E_Refl in ha0.
+    hauto l:on use:Su_Pi_Proj2.
+    have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
+    split.
+    apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
+    move /regularity_sub0 : hSu10 => [i0].
+    hauto l:on use:bind_inst.
+    hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
+    split.
+    by apply : T_App; eauto using T_Conv_E.
+    apply : T_Conv; eauto.
+    apply T_App with (A := A2) (B := B2); eauto.
+    apply : T_Conv_E; eauto.
+    move /E_Symmetric in h01.
+    move /regularity_sub0 : hSu20 => [i0].
+    sfirstorder use:bind_inst.
+  - move => p0 p1 a0 a1 hne0 hne1 doma iha /DJoin.hne_proj_inj /(_ hne0 hne1) [? hja]. subst.
+    move : iha hja; repeat move/[apply].
+    move => [_ iha]. apply ce_neu_neu_helper => // Γ A B.
+    move : iha (hne0) (hne1);repeat move/[apply].
+    move => ih.
+    case : p1.
+    ** move => ha0 ha1.
+       move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0.
+       move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1.
+       move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply].
+       move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
+       split. sauto lq:on.
+       move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
+       have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
+         by eauto using Su_Transitive, Su_Eq.
+       have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
+         by eauto using Su_Transitive, Su_Eq.
+       exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive.
+       repeat split => //=.
+       hauto l:on use:Su_Sig_Proj1, Su_Transitive.
+       apply T_Proj1 with (B := B2); eauto using T_Conv_E.
+       apply T_Proj1 with (B := B2); eauto using T_Conv_E.
+    ** move => ha0  ha1.
+       move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0.
+       move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1.
+       move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply].
+       move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
+       split. sauto lq:on.
+       move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
+       have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
+         by eauto using Su_Transitive, Su_Eq.
+       have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
+         by eauto using Su_Transitive, Su_Eq.
+       have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1.
+       have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1.
+       have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+       have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+       have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2
+         by sauto lq:on use:coqeq_sound_mutual.
+       exists (subst_PTm (scons (PProj PL a0) VarPTm) B2).
+       repeat split.
+       *** apply : Su_Transitive; eauto.
+           have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2
+             by qauto use:regularity, E_Refl.
+           sfirstorder use:Su_Sig_Proj2.
+       *** apply : Su_Transitive; eauto.
+           sfirstorder use:Su_Sig_Proj2.
+       *** eauto using T_Proj2.
+       *** apply : T_Conv.
+           apply : T_Proj2; eauto.
+           move /E_Symmetric in haE.
+           move /regularity_sub0 in hSu21.
+           sfirstorder use:bind_inst.
+  - move {hhPairNeu hhAbsNeu}.
+    move => P0 P1 u0 u1 b0 b1 c0 c1 neu0 neu1 domP ihP  domu ihu domb ihb domc ihc /hne_ind_inj.
+    move => /(_ neu0 neu1)  [hjP][hju][hjb]hjc.
+    apply ce_neu_neu_helper => // Γ A B.
+    move /Ind_Inv => [iP0][wtP0][wta0][wtb0][wtc0]hSu0.
+    move /Ind_Inv => [iP1][wtP1][wta1][wtb1][wtc1]hSu1.
+    have {}iha : u0 ∼ u1 by qauto l:on.
+    have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia.
+    move : T_Univ_Raise wtP0; repeat move/[apply]. move => wtP0.
+    move : T_Univ_Raise wtP1; repeat move/[apply]. move => wtP1.
+    have {}ihP : P0 ⇔ P1 by qauto l:on.
+    set Δ := cons _ _ in wtP0, wtP1, wtc0, wtc1.
+    have wfΔ :⊢ Δ by hauto l:on use:wff_mutual.
+    have hPE : Δ ⊢ P0 ≡ P1 ∈ PUniv (max iP0 iP1)
+      by hauto l:on use:coqeq_sound_mutual.
+    have haE : Γ ⊢ u0 ≡ u1 ∈ PNat
+      by hauto l:on use:coqeq_sound_mutual.
+    have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual.
+    have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P0 ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)).
+    eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero.
+    have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P0
+      by eauto using T_Conv_E.
+    have {}ihb : b0 ⇔ b1 by hauto l:on.
+    have hPSig : Γ ⊢ PBind PSig PNat P0 ≡ PBind PSig PNat P1 ∈ PUniv (Nat.max iP0 iP1) by eauto with wt.
+    set T := ren_PTm shift _ in wtc0.
+    have : (cons P0 Δ) ⊢ c1 ∈ T.
+    apply : T_Conv; eauto. apply : ctx_eq_subst_one; eauto with wt.
+    apply : Su_Eq; eauto.
+    subst T. apply : weakening_su; eauto.
+    eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto.
+    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. 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 : neu0 neu1 ihP iha ihb ihc. clear. sauto lq:on.
+    have hEInd : Γ ⊢ PInd P0 u0 b0 c0 ≡ PInd P1 u1 b1 c1 ∈ subst_PTm (scons u0 VarPTm) P0 by hfcrush use:coqeq_sound_mutual.
+    exists (subst_PTm (scons u0 VarPTm) P0).
+    repeat split => //=; eauto with wt.
+    apply : Su_Transitive.
+    apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt.
+    apply : Su_Eq. apply E_Refl. by apply T_Nat'.
+    apply : Su_Eq. apply hPE. by eauto.
+    move : hEInd. clear. hauto l:on use:regularity.
+  - move => a b ha hb.
+    move {hhPairNeu hhAbsNeu}.
+Admitted.
 
 
-  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 a b.
-  elim /Wf_nat.lt_wf_ind.
-  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 : 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 /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 : a b h. move => [:hnfneu].
-  move => a b h.
-  case => fb; case => fa.
-  - split; last by sfirstorder use:hf_not_hne.
-    move {hnfneu}.
-    case : a h fb fa => //=.
-    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp.
-      move => a0 a1 ha0 _ _ Γ A wt0 wt1.
-      move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]].
-      apply : CE_HRed; eauto using rtc_refl.
-      apply CE_AbsAbs.
-      suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto.
-      have ? : n > 0 by sauto solve+:lia.
-      exists (n - 1). split; first by lia.
-      move : (ha0). rewrite /algo_metric.
-      move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har.
-      apply lored_nsteps_abs_inv in hr0, hr1.
-      move : hr0 => [va' [hr00 hr01]].
-      move : hr1 => [vb' [hr10 hr11]]. move {ih}.
-      exists i,j,va',vb'. subst.
-      suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
-      repeat split =>//=. sfirstorder.
-      simpl in *; by lia.
-      move /algo_metric_join /DJoin.symmetric : ha0.
-      have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
-      move => /[dup] [[ha00 ha10]] [].
-      move : DJoin.abs_inj; repeat move/[apply].
-      move : DJoin.standardization ha00 ha10; repeat move/[apply].
-      (* this is awful *)
-      move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]].
-      have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va'
-        by hauto lq:on use:@relations.rtc_nsteps.
-      have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb'
-        by hauto lq:on use:@relations.rtc_nsteps.
-      simpl in *.
-      have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst.
-      sfirstorder.
-    + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp.
-      move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
-      have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1)
-        by qauto l:on use:fundamental_theorem, logrel.SemWt_SN.
-      apply : CE_HRed; eauto using rtc_refl.
-      move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0.
-      move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1.
-      move /Sub_Bind_InvR :  (hSu0).
-      move => [i][A2][B2]hE.
-      have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2
-        by eauto using Su_Transitive, Su_Eq.
-      have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2
-        by eauto using Su_Transitive, Su_Eq.
-      have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1.
-      have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1.
-      have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
-      have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
-      move /algo_metric_pair : h sn0 sn1; repeat move/[apply].
-      move => [j][hj][ih0 ih1].
-      have haE : a0 ⇔ a1 by hauto l:on.
-      apply CE_PairPair => //.
-      have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2
-        by hauto l:on use:coqeq_sound_mutual.
-      have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2.
-      apply : T_Conv; eauto.
-      move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2.
-      eapply ih; cycle -1; eauto.
-      apply : T_Conv; eauto.
-      apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2).
-      move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2.
-      move : hE haE. clear.
-      move => h.
-      eapply regularity in h.
-      move : h => [_ [hB _]].
-      eauto using bind_inst.
-    + case : b => //=.
-      * hauto lq:on use:T_AbsBind_Imp.
-      * hauto lq:on rew:off use:T_PairBind_Imp.
-      * move => p1 A1 B1 p0 A0 B0.
-        move /algo_metric_bind.
-        move => [?][j [ih0 [ih1 ih2]]]_ _. subst.
-        move => Γ A hU0 hU1.
-        move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]].
-        move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]].
-        have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1)
-          by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia.
-        have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1)
-          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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1)
-          by hauto lq:on use:T_Univ_Raise solve+:lia.
-        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.
-        have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1)
-          by hauto l:on use:coqeq_sound_mutual.
-        apply : CE_HRed; eauto using rtc_refl.
-        constructor => //.
+(*   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 a b. *)
+(*   elim /Wf_nat.lt_wf_ind. *)
+(*   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 : 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 /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 : a b h. move => [:hnfneu]. *)
+(*   move => a b h. *)
+(*   case => fb; case => fa. *)
+(*   - split; last by sfirstorder use:hf_not_hne. *)
+(*     move {hnfneu}. *)
+(*     case : a h fb fa => //=. *)
+(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp. *)
+(*       move => a0 a1 ha0 _ _ Γ A wt0 wt1. *)
+(*       move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]]. *)
+(*       apply : CE_HRed; eauto using rtc_refl. *)
+(*       apply CE_AbsAbs. *)
+(*       suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto. *)
+(*       have ? : n > 0 by sauto solve+:lia. *)
+(*       exists (n - 1). split; first by lia. *)
+(*       move : (ha0). rewrite /algo_metric. *)
+(*       move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har. *)
+(*       apply lored_nsteps_abs_inv in hr0, hr1. *)
+(*       move : hr0 => [va' [hr00 hr01]]. *)
+(*       move : hr1 => [vb' [hr10 hr11]]. move {ih}. *)
+(*       exists i,j,va',vb'. subst. *)
+(*       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0. *)
+(*       repeat split =>//=. sfirstorder. *)
+(*       simpl in *; by lia. *)
+(*       move /algo_metric_join /DJoin.symmetric : ha0. *)
+(*       have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
+(*       move => /[dup] [[ha00 ha10]] []. *)
+(*       move : DJoin.abs_inj; repeat move/[apply]. *)
+(*       move : DJoin.standardization ha00 ha10; repeat move/[apply]. *)
+(*       (* this is awful *) *)
+(*       move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]]. *)
+(*       have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va' *)
+(*         by hauto lq:on use:@relations.rtc_nsteps. *)
+(*       have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb' *)
+(*         by hauto lq:on use:@relations.rtc_nsteps. *)
+(*       simpl in *. *)
+(*       have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst. *)
+(*       sfirstorder. *)
+(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp. *)
+(*       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1. *)
+(*       have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1) *)
+(*         by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
+(*       apply : CE_HRed; eauto using rtc_refl. *)
+(*       move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0. *)
+(*       move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1. *)
+(*       move /Sub_Bind_InvR :  (hSu0). *)
+(*       move => [i][A2][B2]hE. *)
+(*       have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2 *)
+(*         by eauto using Su_Transitive, Su_Eq. *)
+(*       have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 *)
+(*         by eauto using Su_Transitive, Su_Eq. *)
+(*       have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1. *)
+(*       have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1. *)
+(*       have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
+(*       have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
+(*       move /algo_metric_pair : h sn0 sn1; repeat move/[apply]. *)
+(*       move => [j][hj][ih0 ih1]. *)
+(*       have haE : a0 ⇔ a1 by hauto l:on. *)
+(*       apply CE_PairPair => //. *)
+(*       have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2 *)
+(*         by hauto l:on use:coqeq_sound_mutual. *)
+(*       have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2. *)
+(*       apply : T_Conv; eauto. *)
+(*       move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2. *)
+(*       eapply ih; cycle -1; eauto. *)
+(*       apply : T_Conv; eauto. *)
+(*       apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2). *)
+(*       move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2. *)
+(*       move : hE haE. clear. *)
+(*       move => h. *)
+(*       eapply regularity in h. *)
+(*       move : h => [_ [hB _]]. *)
+(*       eauto using bind_inst. *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on use:T_AbsBind_Imp. *)
+(*       * hauto lq:on rew:off use:T_PairBind_Imp. *)
+(*       * move => p1 A1 B1 p0 A0 B0. *)
+(*         move /algo_metric_bind. *)
+(*         move => [?][j [ih0 [ih1 ih2]]]_ _. subst. *)
+(*         move => Γ A hU0 hU1. *)
+(*         move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]]. *)
+(*         move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]]. *)
+(*         have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1) *)
+(*           by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia. *)
+(*         have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1) *)
+(*           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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1) *)
+(*           by hauto lq:on use:T_Univ_Raise solve+:lia. *)
+(*         have {}hB1 : (cons A1 Γ) ⊢ B1 ∈ PUniv (max i0 i1) *)
+(*           by hauto lq:on use:T_Univ_Raise solve+:lia. *)
 
-        eapply ih; eauto.
-        apply : ctx_eq_subst_one; eauto.
-        eauto using Su_Eq.
-      * move => > /algo_metric_join.
-        hauto lq:on use:DJoin.bind_univ_noconf.
-      * move => > /algo_metric_join.
-        hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin.
-      * move => > /algo_metric_join.
-        clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv.
-      * move => > /algo_metric_join. clear.
-        hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv.
-    + case : b => //=.
-      * hauto lq:on use:T_AbsUniv_Imp.
-      * hauto lq:on use:T_PairUniv_Imp.
-      * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric.
-      * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj.
-        subst.
-        hauto l:on.
-      * move => > /algo_metric_join.
-        hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin.
-      * move => > /algo_metric_join.
-        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv.
-      * move => > /algo_metric_join.
-        clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv.
-    + case : b => //=.
-      * qauto l:on use:T_AbsNat_Imp.
-      * qauto l:on use:T_PairNat_Imp.
-      * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf.
-      * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf.
-      * hauto l:on.
-      * move /algo_metric_join.
-        hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv.
-    (* Zero *)
-    + case : b => //=.
-      * hauto lq:on rew:off use:T_AbsZero_Imp.
-      * hauto lq: on use: T_PairZero_Imp.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv.
-      * hauto l:on.
-      * move =>> /algo_metric_join.
-        hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv.
-    (* Suc *)
-    + case : b => //=.
-      * hauto lq:on rew:off use:T_AbsSuc_Imp.
-      * hauto lq:on use:T_PairSuc_Imp.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv.
-      * move => > /algo_metric_join.
-        hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv.
-      * move => a0 a1 h _ _.
-        move /algo_metric_suc : h => [j [h4 h5]].
-        move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3].
-        move : ih h4 h5;do!move/[apply].
-        move => [ih _].
-        move : ih h0 h2;do!move/[apply].
-        move => h. apply : CE_HRed; eauto using rtc_refl.
-        by constructor.
-  - move : a b h fb fa. abstract : hnfneu.
-    move => + b.
-    case : b => //=.
-    (* NeuAbs *)
-    + move => a u halg _ nea. split => // Γ A hu /[dup] ha.
-      move /Abs_Inv => [A0][B0][_]hSu.
-      move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE.
-      have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
-      have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E.
-      have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i
-        by hauto l:on use:regularity.
-      have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j
-          by qauto l:on use:Bind_Inv.
-      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.
-      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.
-      apply : CE_HRed; eauto using rtc_refl.
-      eapply CE_NeuAbs => //=.
-      eapply ih; eauto.
-    (* NeuPair *)
-    + move => a0 b0 u halg _ neu.
-      split => // Γ A hu /[dup] wt.
-      move /Pair_Inv => [A0][B0][ha0][hb0]hU.
-      move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE.
-      have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on.
-      have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
-      move /Pair_Sig_Inv : wt => [{}ha0 {}hb0].
-      have /T_Proj1 huL := hu.
-      have /T_Proj2 {hu} huR := hu.
-      move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]].
-      have heL : PProj PL u ⇔ a0 by hauto l:on.
-      eapply CE_HRed; eauto using rtc_refl.
-      apply CE_NeuPair; eauto.
-      eapply ih; eauto.
-      apply : T_Conv; eauto.
-      have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i
-        by hauto l:on use:regularity.
-      have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by
-        hauto l:on use:coqeq_sound_mutual.
-      hauto l:on use:bind_inst.
-    (* NeuBind: Impossible *)
-    + move => b p p0 a /algo_metric_join h _ h0. exfalso.
-      move : h h0. clear.
-      move /Sub.FromJoin.
-      hauto l:on use:Sub.hne_bind_noconf.
-    (* NeuUniv: Impossible *)
-    + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join.
-    + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join.
-    (* Zero *)
-    + case => //=.
-      * move => i /algo_metric_join. clear.
-        hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv.
-      * move => >/algo_metric_join. clear.
-        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.
-      * move => >/algo_metric_join. clear.
-        move => h _ h2. exfalso.
-        hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv.
-    (* Suc *)
-    + move => a0.
-      case => //=; move => >/algo_metric_join; clear.
-      * 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.
-      * 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 Γ 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.
+(*         have ihA : A0 ⇔ A1 by hauto l:on. *)
+(*         have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1) *)
+(*           by hauto l:on use:coqeq_sound_mutual. *)
+(*         apply : CE_HRed; eauto using rtc_refl. *)
+(*         constructor => //. *)
 
-    case : a h fb fa => //=.
-    + case : b => //=; move => > /algo_metric_join.
-      * 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.
-      * clear => ? ? _. exfalso.
-        hauto q:on use:REReds.var_inv, REReds.hne_ind_inv.
-    + case : b => //=;
-                   lazymatch goal with
-                   | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac
-                   | _ => move => > /algo_metric_join
-                   end.
-      * clear => *. exfalso.
-        hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv.
-      (* real case *)
-      * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB.
-        move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0.
-        move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1.
-        move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply].
-        move => [j [hj [hal0 hal1]]].
-        have /ih := hal0.
-        move /(_ hj).
-        move => [_ ihb].
-        move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply].
-        move => [hb01][C][hT0][hT1][hT2]hT3.
-        move /Sub_Bind_InvL : (hT0).
-        move => [i][A2][B2]hE.
-        have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by
-          eauto using Su_Eq, Su_Transitive.
-        have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by
-          eauto using Su_Eq, Su_Transitive.
-        have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
-        have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1.
-        have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
-        have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
-        move : ih (hj) hal1. repeat move/[apply].
-        move => [ih _].
-        move : ih (ha0') (ha1'); repeat move/[apply].
-        move => iha.
-        split. qblast.
-        exists (subst_PTm (scons a0 VarPTm) B2).
-        split.
-        apply : Su_Transitive; eauto.
-        move /E_Refl in ha0.
-        hauto l:on use:Su_Pi_Proj2.
-        have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual.
-        split.
-        apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2).
-        move /regularity_sub0 : hSu10 => [i0].
-        hauto l:on use:bind_inst.
-        hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl.
-        split.
-        by apply : T_App; eauto using T_Conv_E.
-        apply : T_Conv; eauto.
-        apply T_App with (A := A2) (B := B2); eauto.
-        apply : T_Conv_E; eauto.
-        move /E_Symmetric in h01.
-        move /regularity_sub0 : hSu20 => [i0].
-        sfirstorder use:bind_inst.
-      * clear => ? ? ?. exfalso.
-        hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv.
-      * clear => ? ? ?. exfalso.
-        hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv.
-    + case : b => //=.
-      * move => i p h /algo_metric_join. clear => ? _ ?. exfalso.
-        hauto l:on use:REReds.hne_proj_inv, REReds.var_inv.
-      * move => > /algo_metric_join. clear => ? ? ?. exfalso.
-        hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv.
-      (* real case *)
-      * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0.
-        move : ha (hne0) (hne1); repeat move/[apply].
-        move => [? [j []]]. subst.
-        move : ih; repeat move/[apply].
-        move => [_ ih].
-        case : p1.
-        ** move => Γ A B ha0 ha1.
-           move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0.
-           move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1.
-           move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply].
-           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
-           split. sauto lq:on.
-           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
-           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
-             by eauto using Su_Transitive, Su_Eq.
-           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
-             by eauto using Su_Transitive, Su_Eq.
-           exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive.
-           repeat split => //=.
-           hauto l:on use:Su_Sig_Proj1, Su_Transitive.
-           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
-           apply T_Proj1 with (B := B2); eauto using T_Conv_E.
-        ** move => Γ A B  ha0  ha1.
-           move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0.
-           move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1.
-           move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply].
-           move => [ha [C [hS0 [hS1 [wta0 wta1]]]]].
-           split. sauto lq:on.
-           move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2.
-           have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0
-             by eauto using Su_Transitive, Su_Eq.
-           have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1
-             by eauto using Su_Transitive, Su_Eq.
-           have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1.
-           have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1.
-           have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
-           have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
-           have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2
-             by sauto lq:on use:coqeq_sound_mutual.
-           exists (subst_PTm (scons (PProj PL a0) VarPTm) B2).
-           repeat split.
-           *** apply : Su_Transitive; eauto.
-               have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2
-                 by qauto use:regularity, E_Refl.
-               sfirstorder use:Su_Sig_Proj2.
-           *** apply : Su_Transitive; eauto.
-               sfirstorder use:Su_Sig_Proj2.
-           *** eauto using T_Proj2.
-           *** apply : T_Conv.
-               apply : T_Proj2; eauto.
-               move /E_Symmetric in haE.
-               move /regularity_sub0 in hSu21.
-               sfirstorder use:bind_inst.
-      * move => > /algo_metric_join; clear => ? ? ?. exfalso.
-        hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv.
-    (* ind ind case *)
-    + move => P a0 b0 c0.
-      case : b => //=;
-                   lazymatch goal with
-                   | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac
-                   | _ => move => > /algo_metric_join; clear => *; exfalso
-                   end.
-      * 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.
-      * 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.
-        move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0.
-        move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1.
-        have {}iha : a0 ∼ a1 by qauto l:on.
-        have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia.
-        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 Δ := 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.
-        have haE : Γ ⊢ a0 ≡ a1 ∈ PNat
-          by hauto l:on use:coqeq_sound_mutual.
-        have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual.
-        have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)).
-        eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero.
-        have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P
-          by eauto using T_Conv_E.
-        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 : (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.
-        eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto.
-        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. 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.
-        have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual.
-        exists (subst_PTm (scons a0 VarPTm) P).
-        repeat split => //=; eauto with wt.
-        apply : Su_Transitive.
-        apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt.
-        apply : Su_Eq. apply E_Refl. by apply T_Nat'.
-        apply : Su_Eq. apply hPE. by eauto.
-        move : hEInd. clear. hauto l:on use:regularity.
-Qed.
+(*         eapply ih; eauto. *)
+(*         apply : ctx_eq_subst_one; eauto. *)
+(*         eauto using Su_Eq. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on use:DJoin.bind_univ_noconf. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin. *)
+(*       * move => > /algo_metric_join. *)
+(*         clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv. *)
+(*       * move => > /algo_metric_join. clear. *)
+(*         hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv. *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on use:T_AbsUniv_Imp. *)
+(*       * hauto lq:on use:T_PairUniv_Imp. *)
+(*       * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric. *)
+(*       * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj. *)
+(*         subst. *)
+(*         hauto l:on. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin. *)
+(*       * move => > /algo_metric_join. *)
+(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv. *)
+(*     + case : b => //=. *)
+(*       * qauto l:on use:T_AbsNat_Imp. *)
+(*       * qauto l:on use:T_PairNat_Imp. *)
+(*       * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf. *)
+(*       * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf. *)
+(*       * hauto l:on. *)
+(*       * move /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv. *)
+(*     (* Zero *) *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on rew:off use:T_AbsZero_Imp. *)
+(*       * hauto lq: on use: T_PairZero_Imp. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv. *)
+(*       * hauto l:on. *)
+(*       * move =>> /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv. *)
+(*     (* Suc *) *)
+(*     + case : b => //=. *)
+(*       * hauto lq:on rew:off use:T_AbsSuc_Imp. *)
+(*       * hauto lq:on use:T_PairSuc_Imp. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv. *)
+(*       * move => > /algo_metric_join. *)
+(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv. *)
+(*       * move => a0 a1 h _ _. *)
+(*         move /algo_metric_suc : h => [j [h4 h5]]. *)
+(*         move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3]. *)
+(*         move : ih h4 h5;do!move/[apply]. *)
+(*         move => [ih _]. *)
+(*         move : ih h0 h2;do!move/[apply]. *)
+(*         move => h. apply : CE_HRed; eauto using rtc_refl. *)
+(*         by constructor. *)
+(*   - move : a b h fb fa. abstract : hnfneu. *)
+(*     move => + b. *)
+(*     case : b => //=. *)
+(*     (* NeuAbs *) *)
+(*     + move => a u halg _ nea. split => // Γ A hu /[dup] ha. *)
+(*       move /Abs_Inv => [A0][B0][_]hSu. *)
+(*       move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE. *)
+(*       have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
+(*       have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
+(*       have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i *)
+(*         by hauto l:on use:regularity. *)
+(*       have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j *)
+(*           by qauto l:on use:Bind_Inv. *)
+(*       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. *)
+(*       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. *)
+(*       apply : CE_HRed; eauto using rtc_refl. *)
+(*       eapply CE_NeuAbs => //=. *)
+(*       eapply ih; eauto. *)
+(*     (* NeuPair *) *)
+(*     + move => a0 b0 u halg _ neu. *)
+(*       split => // Γ A hu /[dup] wt. *)
+(*       move /Pair_Inv => [A0][B0][ha0][hb0]hU. *)
+(*       move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE. *)
+(*       have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on. *)
+(*       have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
+(*       move /Pair_Sig_Inv : wt => [{}ha0 {}hb0]. *)
+(*       have /T_Proj1 huL := hu. *)
+(*       have /T_Proj2 {hu} huR := hu. *)
+(*       move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]]. *)
+(*       have heL : PProj PL u ⇔ a0 by hauto l:on. *)
+(*       eapply CE_HRed; eauto using rtc_refl. *)
+(*       apply CE_NeuPair; eauto. *)
+(*       eapply ih; eauto. *)
+(*       apply : T_Conv; eauto. *)
+(*       have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i *)
+(*         by hauto l:on use:regularity. *)
+(*       have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by *)
+(*         hauto l:on use:coqeq_sound_mutual. *)
+(*       hauto l:on use:bind_inst. *)
+(*     (* NeuBind: Impossible *) *)
+(*     + move => b p p0 a /algo_metric_join h _ h0. exfalso. *)
+(*       move : h h0. clear. *)
+(*       move /Sub.FromJoin. *)
+(*       hauto l:on use:Sub.hne_bind_noconf. *)
+(*     (* NeuUniv: Impossible *) *)
+(*     + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join. *)
+(*     + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join. *)
+(*     (* Zero *) *)
+(*     + case => //=. *)
+(*       * move => i /algo_metric_join. clear. *)
+(*         hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv. *)
+(*       * move => >/algo_metric_join. clear. *)
+(*         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. *)
+(*       * move => >/algo_metric_join. clear. *)
+(*         move => h _ h2. exfalso. *)
+(*         hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv. *)
+(*     (* Suc *) *)
+(*     + move => a0. *)
+(*       case => //=; move => >/algo_metric_join; clear. *)
+(*       * 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. *)
+(*       * 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 Γ 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 => 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. *)
+(*       * clear => ? ? _. exfalso. *)
+(*         hauto q:on use:REReds.var_inv, REReds.hne_ind_inv. *)
+(*     + case : b => //=; *)
+(*                    lazymatch goal with *)
+(*                    | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac *)
+(*                    | _ => move => > /algo_metric_join *)
+(*                    end. *)
+(*       * clear => *. exfalso. *)
+(*         hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv. *)
+(*       (* real case *) *)
+(*       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB. *)
+(*         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0. *)
+(*         move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1. *)
+(*         move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply]. *)
+(*         move => [j [hj [hal0 hal1]]]. *)
+(*         have /ih := hal0. *)
+(*         move /(_ hj). *)
+(*         move => [_ ihb]. *)
+(*         move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply]. *)
+(*         move => [hb01][C][hT0][hT1][hT2]hT3. *)
+(*         move /Sub_Bind_InvL : (hT0). *)
+(*         move => [i][A2][B2]hE. *)
+(*         have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by *)
+(*           eauto using Su_Eq, Su_Transitive. *)
+(*         have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by *)
+(*           eauto using Su_Eq, Su_Transitive. *)
+(*         have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
+(*         have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
+(*         have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
+(*         have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
+(*         move : ih (hj) hal1. repeat move/[apply]. *)
+(*         move => [ih _]. *)
+(*         move : ih (ha0') (ha1'); repeat move/[apply]. *)
+(*         move => iha. *)
+(*         split. qblast. *)
+(*         exists (subst_PTm (scons a0 VarPTm) B2). *)
+(*         split. *)
+(*         apply : Su_Transitive; eauto. *)
+(*         move /E_Refl in ha0. *)
+(*         hauto l:on use:Su_Pi_Proj2. *)
+(*         have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual. *)
+(*         split. *)
+(*         apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2). *)
+(*         move /regularity_sub0 : hSu10 => [i0]. *)
+(*         hauto l:on use:bind_inst. *)
+(*         hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl. *)
+(*         split. *)
+(*         by apply : T_App; eauto using T_Conv_E. *)
+(*         apply : T_Conv; eauto. *)
+(*         apply T_App with (A := A2) (B := B2); eauto. *)
+(*         apply : T_Conv_E; eauto. *)
+(*         move /E_Symmetric in h01. *)
+(*         move /regularity_sub0 : hSu20 => [i0]. *)
+(*         sfirstorder use:bind_inst. *)
+(*       * clear => ? ? ?. exfalso. *)
+(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv. *)
+(*       * clear => ? ? ?. exfalso. *)
+(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv. *)
+(*     + case : b => //=. *)
+(*       * move => i p h /algo_metric_join. clear => ? _ ?. exfalso. *)
+(*         hauto l:on use:REReds.hne_proj_inv, REReds.var_inv. *)
+(*       * move => > /algo_metric_join. clear => ? ? ?. exfalso. *)
+(*         hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv. *)
+(*       (* real case *) *)
+(*       * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0. *)
+(*         move : ha (hne0) (hne1); repeat move/[apply]. *)
+(*         move => [? [j []]]. subst. *)
+(*         move : ih; repeat move/[apply]. *)
+(*         move => [_ ih]. *)
+(*         case : p1. *)
+(*         ** move => Γ A B ha0 ha1. *)
+(*            move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
+(*            move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
+(*            move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply]. *)
+(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
+(*            split. sauto lq:on. *)
+(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
+(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive. *)
+(*            repeat split => //=. *)
+(*            hauto l:on use:Su_Sig_Proj1, Su_Transitive. *)
+(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
+(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
+(*         ** move => Γ A B  ha0  ha1. *)
+(*            move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
+(*            move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
+(*            move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply]. *)
+(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
+(*            split. sauto lq:on. *)
+(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
+(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
+(*              by eauto using Su_Transitive, Su_Eq. *)
+(*            have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1. *)
+(*            have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1. *)
+(*            have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
+(*            have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
+(*            have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2 *)
+(*              by sauto lq:on use:coqeq_sound_mutual. *)
+(*            exists (subst_PTm (scons (PProj PL a0) VarPTm) B2). *)
+(*            repeat split. *)
+(*            *** apply : Su_Transitive; eauto. *)
+(*                have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2 *)
+(*                  by qauto use:regularity, E_Refl. *)
+(*                sfirstorder use:Su_Sig_Proj2. *)
+(*            *** apply : Su_Transitive; eauto. *)
+(*                sfirstorder use:Su_Sig_Proj2. *)
+(*            *** eauto using T_Proj2. *)
+(*            *** apply : T_Conv. *)
+(*                apply : T_Proj2; eauto. *)
+(*                move /E_Symmetric in haE. *)
+(*                move /regularity_sub0 in hSu21. *)
+(*                sfirstorder use:bind_inst. *)
+(*       * move => > /algo_metric_join; clear => ? ? ?. exfalso. *)
+(*         hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv. *)
+(*     (* ind ind case *) *)
+(*     + move => P a0 b0 c0. *)
+(*       case : b => //=; *)
+(*                    lazymatch goal with *)
+(*                    | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac *)
+(*                    | _ => move => > /algo_metric_join; clear => *; exfalso *)
+(*                    end. *)
+(*       * 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. *)
+(*       * 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. *)
+(*         move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0. *)
+(*         move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1. *)
+(*         have {}iha : a0 ∼ a1 by qauto l:on. *)
+(*         have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia. *)
+(*         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 Δ := 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. *)
+(*         have haE : Γ ⊢ a0 ≡ a1 ∈ PNat *)
+(*           by hauto l:on use:coqeq_sound_mutual. *)
+(*         have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual. *)
+(*         have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)). *)
+(*         eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero. *)
+(*         have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P *)
+(*           by eauto using T_Conv_E. *)
+(*         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 : (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. *)
+(*         eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto. *)
+(*         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. 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. *)
+(*         have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual. *)
+(*         exists (subst_PTm (scons a0 VarPTm) P). *)
+(*         repeat split => //=; eauto with wt. *)
+(*         apply : Su_Transitive. *)
+(*         apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt. *)
+(*         apply : Su_Eq. apply E_Refl. by apply T_Nat'. *)
+(*         apply : Su_Eq. apply hPE. by eauto. *)
+(*         move : hEInd. clear. hauto l:on use:regularity. *)
+(* Qed. *)
 
 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 Γ (a b A : PTm) :
-  Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
-Proof.
-  move => h.
-  suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity.
-  eapply fundamental_theorem in h.
-  move /logrel.SemEq_SN_Join : h.
-  move => h.
-  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.
-  move => [va][vb][hva][hvb][nva][nvb]hj.
-  move /relations.rtc_nsteps : hva => [i hva].
-  move /relations.rtc_nsteps : hvb => [j hvb].
-  exists (i + j + size_PTm va + size_PTm vb).
-  hauto lq:on solve+:lia.
-Qed.
+(* Lemma coqeq_complete Γ (a b A : PTm) : *)
+(*   Γ ⊢ a ≡ b ∈ A -> a ⇔ b. *)
+(* Proof. *)
+(*   move => h. *)
+(*   suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity. *)
+(*   eapply fundamental_theorem in h. *)
+(*   move /logrel.SemEq_SN_Join : h. *)
+(*   move => h. *)
+(*   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. *)
+(*   move => [va][vb][hva][hvb][nva][nvb]hj. *)
+(*   move /relations.rtc_nsteps : hva => [i hva]. *)
+(*   move /relations.rtc_nsteps : hvb => [j hvb]. *)
+(*   exists (i + j + size_PTm va + size_PTm vb). *)
+(*   hauto lq:on solve+:lia. *)
+(* Qed. *)
 
 
 Reserved Notation "a ≪ b" (at level 70).
diff --git a/theories/common.v b/theories/common.v
index 18f4dfb..39713ba 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -284,8 +284,8 @@ Inductive algo_dom : PTm -> PTm -> Prop :=
   algo_dom (PInd P0 u0 b0 c0) (PInd P1 u1 b1 c1)
 
 | A_Conf a b :
-  HRed.nf a ->
-  HRed.nf b ->
+  ishf a \/ ishne a ->
+  ishf b \/ ishne b ->
   tm_conf a b ->
   algo_dom a b
 
@@ -376,8 +376,8 @@ Inductive salgo_dom : PTm -> PTm -> Prop :=
   salgo_dom a b
 
 | S_Conf a b :
-  HRed.nf a ->
-  HRed.nf b ->
+  ishf a \/ ishne a ->
+  ishf b \/ ishne b ->
   stm_conf a b ->
   salgo_dom a b
 
@@ -417,7 +417,7 @@ Proof. elim : a b => //=; hauto l:on inv:HRed.R. Qed.
 
 Ltac2 destruct_salgo () :=
   lazy_match! goal with
-  | [h : is_true (ishne ?a) |- is_true (stm_conf ?a _) ] =>
+  | [_ : is_true (ishne ?a) |- is_true (stm_conf ?a _) ] =>
       if Constr.is_var a then destruct $a; ltac1:(done) else ()
   | [|- is_true (stm_conf _ _)] =>
       unfold stm_conf; ltac1:(done)
diff --git a/theories/executable.v b/theories/executable.v
index 238fe45..e62e3c3 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -226,7 +226,7 @@ Ltac check_sub_triv :=
   intros;subst;
   lazymatch goal with
   (* | [h : algo_dom (VarPTm _) (PAbs _) |- _] => idtac *)
-  | [h : salgo_dom _ _ |- _] => try (inversion h; subst => //=; ltac2:(Control.enter destruct_algo))
+  | [_ : salgo_dom _ _ |- _] => try (inversion h; subst => //=; ltac2:(Control.enter destruct_algo))
   | _ => idtac
   end.
 
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index d51af9d..4b76638 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -246,9 +246,6 @@ Proof.
     move => /negP h0.
     eapply check_equal_complete in h0.
     apply h0. by constructor.
-  - move => a b i0 i1 j. simp ce_prop.
-    move => _ h. inversion h; subst => //=.
-    hauto lq:on inv:CoqEq_Neu unfold:stm_conf.
   - move => a b s ihs. simp ce_prop.
     move => h0 h1. apply ihs =>//.
     have [? ?] : HRed.nf a /\ HRed.nf b by hauto l:on use:salgo_dom_no_hred.

From 9d68e5d6c93b17a033abea7ea0439af364c15301 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 23:22:09 -0400
Subject: [PATCH 195/210] Finish the conf case

---
 theories/algorithmic.v | 63 +++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 62 insertions(+), 1 deletion(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index eb93263..c8a4799 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -448,6 +448,14 @@ Proof.
   hauto lq:on ctrs:rtc, CoqEq_R inv:CoqEq_R.
 Qed.
 
+Lemma CE_HRedR (a b b' : PTm) :
+  HRed.R b b' ->
+  a ⇔ b' ->
+  a ⇔ b.
+Proof.
+  hauto lq:on ctrs:rtc, CoqEq_R inv:CoqEq_R.
+Qed.
+
 Lemma CE_Nf a b :
   a ↔ b -> a ⇔ b.
 Proof. hauto l:on ctrs:rtc. Qed.
@@ -1316,7 +1324,7 @@ Proof. hauto q:on use:REReds.hne_ind_inv. Qed.
 Lemma coqeq_complete' :
   (forall a b, algo_dom a b -> DJoin.R 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)) /\
     (forall a b, algo_dom_r a b -> DJoin.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⇔ b).
-  move => [:hhPairNeu hhAbsNeu].
+  move => [:hConfNeuNf hhPairNeu hhAbsNeu].
   apply algo_dom_mutual.
   - move => a b h ih hj. split => //.
     move => Γ A. move : T_Abs_Inv; repeat move/[apply].
@@ -1516,6 +1524,59 @@ Lemma coqeq_complete' :
     move : hEInd. clear. hauto l:on use:regularity.
   - move => a b ha hb.
     move {hhPairNeu hhAbsNeu}.
+    case : hb; case : ha.
+    + move {hConfNeuNf}.
+      move => h0 h1 h2 h3. split; last by sfirstorder use:hf_not_hne.
+      move : h0 h1 h2 h3.
+      case : b; case : a => //= *; try by (exfalso; eauto 2 using T_AbsPair_Imp, T_AbsUniv_Imp, T_AbsBind_Imp, T_AbsNat_Imp, T_AbsZero_Imp, T_AbsSuc_Imp,  T_PairUniv_Imp, T_PairBind_Imp, T_PairNat_Imp, T_PairZero_Imp, T_PairSuc_Imp).
+      sfirstorder use:DJoin.bind_univ_noconf.
+      hauto q:on use:REReds.nat_inv, REReds.bind_inv.
+      hauto q:on use:REReds.zero_inv, REReds.bind_inv.
+      hauto q:on use:REReds.suc_inv, REReds.bind_inv.
+      hauto q:on use:REReds.bind_inv, REReds.univ_inv.
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv.
+      hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv.
+      hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv.
+      hauto lq:on rew:off use:REReds.bind_inv, REReds.nat_inv.
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv.
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv.
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv.
+      hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv.
+      hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv.
+      hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv.
+      hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv.
+      hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv.
+      hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv.
+      hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv.
+      hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv.
+    + abstract : hConfNeuNf a b.
+      move => h0 h1 h2 h3. split; last by sfirstorder use:hf_not_hne.
+      move : h0 h1 h2 h3.
+      case : b; case : a => //=; hauto q:on use:REReds.var_inv, REReds.bind_inv, REReds.hne_app_inv, REReds.hne_proj_inv, REReds.hne_ind_inv, REReds.bind_inv, REReds.nat_inv, REReds.univ_inv, REReds.zero_inv, REReds.suc_inv.
+    + rewrite tm_conf_sym.
+      move => h0 h1 h2 /DJoin.symmetric hb.
+      move : hConfNeuNf h0 h1 h2 hb; repeat move/[apply].
+      qauto l:on use:coqeq_symmetric_mutual.
+    + move => neua neub hconf hj.
+      move {hConfNeuNf}.
+      exfalso.
+      move : neua neub hconf hj.
+      case : b; case : a => //=*; exfalso; hauto q:on use:REReds.var_inv, REReds.bind_inv, REReds.hne_app_inv, REReds.hne_proj_inv, REReds.hne_ind_inv.
+  - sfirstorder.
+  - move {hConfNeuNf hhPairNeu hhAbsNeu}.
+    move => a a' b hr ha iha hj Γ A wta wtb.
+    apply : CE_HRedL; eauto.
+    apply : iha; eauto; last by sfirstorder use:HRed.preservation.
+    apply : DJoin.transitive; eauto.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply DJoin.FromRRed1. by apply HRed.ToRRed.
+  - move {hConfNeuNf hhPairNeu hhAbsNeu}.
+    move => a b b' nfa hr h ih j Γ A wta wtb.
+    apply : CE_HRedR; eauto.
+    apply : ih; eauto; last by eauto using HRed.preservation.
+    apply : DJoin.transitive; eauto.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply DJoin.FromRRed0. by apply HRed.ToRRed.
 Admitted.
 
 

From 181e06ae01fed64158a786da5caac24a16108f78 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Mon, 10 Mar 2025 23:45:19 -0400
Subject: [PATCH 196/210] Finish the completeness proof for equality

---
 theories/algorithmic.v | 570 +++++++----------------------------------
 1 file changed, 87 insertions(+), 483 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index c8a4799..a870f9f 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1345,10 +1345,74 @@ Lemma coqeq_complete' :
     apply DJoin.symmetric. apply DJoin.FromEJoin. eexists. split. apply relations.rtc_once.
     apply ERed.AppEta. apply rtc_refl.
   - hauto q:on use:coqeq_symmetric_mutual, DJoin.symmetric, algo_dom_sym.
-  - move {hhAbsNeu hhPairNeu}.
-    admit.
-  - move {hhAbsNeu}. abstract : hhPairNeu.
-    admit.
+  - move {hhAbsNeu hhPairNeu hConfNeuNf}.
+    move => a0 a1 b0 b1 doma iha domb ihb /DJoin.pair_inj hj.
+    split => //.
+    move => Γ A wt0 wt1.
+    have [] : SN (PPair a0 b0) /\ SN (PPair a1 b1) by hauto lq:on use:logrel.SemWt_SN, fundamental_theorem.
+    move : hj; repeat move/[apply].
+    move => [hja hjb].
+    move /Pair_Inv : wt0 => [A0][B0][ha0][hb0]hSu0.
+    move /Pair_Inv : wt1 => [A1][B1][ha1][hb1]hSu1.
+    move /Sub_Bind_InvR :  (hSu0).
+    move => [i][A2][B2]hE.
+    have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2
+      by eauto using Su_Transitive, Su_Eq.
+    have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2
+      by eauto using Su_Transitive, Su_Eq.
+    have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1.
+    have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1.
+    have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv.
+    have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv.
+    move : iha (ha0A2) (ha1A2) hja; repeat move/[apply].
+    move => h.
+    apply CE_Nf.
+    apply CE_PairPair => //.
+    have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2
+      by hauto l:on use:coqeq_sound_mutual.
+    have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2.
+    apply : T_Conv; eauto.
+    move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2.
+    eapply ihb; cycle -1; eauto.
+    apply : T_Conv; eauto.
+    apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2).
+    move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2.
+    move : hE haE. clear.
+    move => h.
+    eapply regularity in h.
+    move : h => [_ [hB _]].
+    eauto using bind_inst.
+  - abstract : hhPairNeu. move {hConfNeuNf}.
+    move => a0 b0 u neu doma iha domb ihb hj.
+    split => // Γ A /[dup] wt /Pair_Inv
+               [A0][B0][ha0][hb0]hU.
+    move => wtu.
+    move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE.
+    have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on.
+    have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
+    move /Pair_Sig_Inv : wt => [{}ha0 {}hb0].
+    have /T_Proj1 huL := hu.
+    have /T_Proj2 {hu} huR := hu.
+    have heL : a0 ⇔ PProj PL u . apply : iha; eauto.
+    apply : DJoin.transitive; cycle 2. apply DJoin.ProjCong; eauto.
+    apply : N_Exp; eauto. apply N_ProjPairL.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply DJoin.FromRRed1. apply RRed.ProjPair.
+    eapply CE_HRed; eauto using rtc_refl.
+    apply CE_PairNeu; eauto.
+    eapply ihb; eauto.
+    apply : DJoin.transitive; cycle 2. apply DJoin.ProjCong; eauto.
+    apply : N_Exp; eauto. apply N_ProjPairR.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply DJoin.FromRRed1. apply RRed.ProjPair.
+    apply : T_Conv; eauto.
+    have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i
+      by hauto l:on use:regularity.
+    have /E_Symmetric : Γ ⊢ a0 ≡ PProj PL u ∈ A2 by
+      hauto l:on use:coqeq_sound_mutual.
+    hauto l:on use:bind_inst.
   - move {hhAbsNeu}.
     move => a0 a1 u hu  ha iha hb ihb /DJoin.symmetric hj. split => // *.
     eapply coqeq_symmetric_mutual.
@@ -1577,490 +1641,30 @@ Lemma coqeq_complete' :
     apply : DJoin.transitive; eauto.
     hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
     apply DJoin.FromRRed0. by apply HRed.ToRRed.
-Admitted.
-
-
-
-(*   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 a b. *)
-(*   elim /Wf_nat.lt_wf_ind. *)
-(*   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 : 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 /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 : a b h. move => [:hnfneu]. *)
-(*   move => a b h. *)
-(*   case => fb; case => fa. *)
-(*   - split; last by sfirstorder use:hf_not_hne. *)
-(*     move {hnfneu}. *)
-(*     case : a h fb fa => //=. *)
-(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_AbsBind_Imp, T_AbsUniv_Imp, T_AbsZero_Imp, T_AbsSuc_Imp, T_AbsNat_Imp. *)
-(*       move => a0 a1 ha0 _ _ Γ A wt0 wt1. *)
-(*       move : T_Abs_Inv wt0 wt1; repeat move/[apply]. move  => [Δ [V [wt1 wt0]]]. *)
-(*       apply : CE_HRed; eauto using rtc_refl. *)
-(*       apply CE_AbsAbs. *)
-(*       suff [l [h0 h1]] : exists l, l < n /\ algo_metric l a1 a0 by eapply ih; eauto. *)
-(*       have ? : n > 0 by sauto solve+:lia. *)
-(*       exists (n - 1). split; first by lia. *)
-(*       move : (ha0). rewrite /algo_metric. *)
-(*       move => [i][j][va][vb][hr0][hr1][nfva][nfvb][[v [hr0' hr1']]] har. *)
-(*       apply lored_nsteps_abs_inv in hr0, hr1. *)
-(*       move : hr0 => [va' [hr00 hr01]]. *)
-(*       move : hr1 => [vb' [hr10 hr11]]. move {ih}. *)
-(*       exists i,j,va',vb'. subst. *)
-(*       suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0. *)
-(*       repeat split =>//=. sfirstorder. *)
-(*       simpl in *; by lia. *)
-(*       move /algo_metric_join /DJoin.symmetric : ha0. *)
-(*       have : SN a0 /\ SN a1 by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
-(*       move => /[dup] [[ha00 ha10]] []. *)
-(*       move : DJoin.abs_inj; repeat move/[apply]. *)
-(*       move : DJoin.standardization ha00 ha10; repeat move/[apply]. *)
-(*       (* this is awful *) *)
-(*       move => [vb][va][h' [h'' [h''' [h'''' h'''''']]]]. *)
-(*       have /LoReds.ToRReds {}hr00 : rtc LoRed.R a1 va' *)
-(*         by hauto lq:on use:@relations.rtc_nsteps. *)
-(*       have /LoReds.ToRReds {}hr10 : rtc LoRed.R a0 vb' *)
-(*         by hauto lq:on use:@relations.rtc_nsteps. *)
-(*       simpl in *. *)
-(*       have [*] : va' = va /\ vb' = vb by eauto using red_uniquenf. subst. *)
-(*       sfirstorder. *)
-(*     + case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp, T_PairNat_Imp, T_PairSuc_Imp, T_PairZero_Imp. *)
-(*       move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1. *)
-(*       have [sn0 sn1] : SN (PPair a0 b0) /\ SN (PPair a1 b1) *)
-(*         by qauto l:on use:fundamental_theorem, logrel.SemWt_SN. *)
-(*       apply : CE_HRed; eauto using rtc_refl. *)
-(*       move /Pair_Inv : hu0 => [A0][B0][ha0][hb0]hSu0. *)
-(*       move /Pair_Inv : hu1 => [A1][B1][ha1][hb1]hSu1. *)
-(*       move /Sub_Bind_InvR :  (hSu0). *)
-(*       move => [i][A2][B2]hE. *)
-(*       have hSu12 : Γ ⊢ PBind PSig A1 B1 ≲ PBind PSig A2 B2 *)
-(*         by eauto using Su_Transitive, Su_Eq. *)
-(*       have hSu02 : Γ ⊢ PBind PSig A0 B0 ≲ PBind PSig A2 B2 *)
-(*         by eauto using Su_Transitive, Su_Eq. *)
-(*       have hA02 : Γ ⊢ A0 ≲ A2 by eauto using Su_Sig_Proj1. *)
-(*       have hA12 : Γ ⊢ A1 ≲ A2 by eauto using Su_Sig_Proj1. *)
-(*       have ha0A2 : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
-(*       have ha1A2 : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
-(*       move /algo_metric_pair : h sn0 sn1; repeat move/[apply]. *)
-(*       move => [j][hj][ih0 ih1]. *)
-(*       have haE : a0 ⇔ a1 by hauto l:on. *)
-(*       apply CE_PairPair => //. *)
-(*       have {}haE : Γ ⊢ a0 ≡ a1 ∈ A2 *)
-(*         by hauto l:on use:coqeq_sound_mutual. *)
-(*       have {}hb1 : Γ ⊢ b1 ∈ subst_PTm (scons a1 VarPTm) B2. *)
-(*       apply : T_Conv; eauto. *)
-(*       move /E_Refl in ha1. hauto l:on use:Su_Sig_Proj2. *)
-(*       eapply ih; cycle -1; eauto. *)
-(*       apply : T_Conv; eauto. *)
-(*       apply Su_Transitive with (B := subst_PTm (scons a0 VarPTm) B2). *)
-(*       move /E_Refl in ha0. hauto l:on use:Su_Sig_Proj2. *)
-(*       move : hE haE. clear. *)
-(*       move => h. *)
-(*       eapply regularity in h. *)
-(*       move : h => [_ [hB _]]. *)
-(*       eauto using bind_inst. *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on use:T_AbsBind_Imp. *)
-(*       * hauto lq:on rew:off use:T_PairBind_Imp. *)
-(*       * move => p1 A1 B1 p0 A0 B0. *)
-(*         move /algo_metric_bind. *)
-(*         move => [?][j [ih0 [ih1 ih2]]]_ _. subst. *)
-(*         move => Γ A hU0 hU1. *)
-(*         move /Bind_Inv : hU0 => [i0 [hA0 [hB0 hS0]]]. *)
-(*         move /Bind_Inv : hU1 => [i1 [hA1 [hB1 hS1]]]. *)
-(*         have ? : Γ ⊢ PUniv i0 ≲ PUniv (max i0 i1) *)
-(*           by hauto lq:on rew:off use:Su_Univ, wff_mutual solve+:lia. *)
-(*         have ? : Γ ⊢ PUniv i1 ≲ PUniv (max i0 i1) *)
-(*           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 : (cons A0 Γ) ⊢ B0 ∈ PUniv (max i0 i1) *)
-(*           by hauto lq:on use:T_Univ_Raise solve+:lia. *)
-(*         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. *)
-(*         have hAE : Γ ⊢ A0 ≡ A1 ∈ PUniv (Nat.max i0 i1) *)
-(*           by hauto l:on use:coqeq_sound_mutual. *)
-(*         apply : CE_HRed; eauto using rtc_refl. *)
-(*         constructor => //. *)
-
-(*         eapply ih; eauto. *)
-(*         apply : ctx_eq_subst_one; eauto. *)
-(*         eauto using Su_Eq. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on use:DJoin.bind_univ_noconf. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on use:Sub.nat_bind_noconf, Sub.FromJoin. *)
-(*       * move => > /algo_metric_join. *)
-(*         clear. hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv. *)
-(*       * move => > /algo_metric_join. clear. *)
-(*         hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv. *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on use:T_AbsUniv_Imp. *)
-(*       * hauto lq:on use:T_PairUniv_Imp. *)
-(*       * qauto l:on use:algo_metric_join, DJoin.bind_univ_noconf, DJoin.symmetric. *)
-(*       * move => i j /algo_metric_join /DJoin.univ_inj ? _ _ Γ A hi hj. *)
-(*         subst. *)
-(*         hauto l:on. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on use:Sub.nat_univ_noconf, Sub.FromJoin. *)
-(*       * move => > /algo_metric_join. *)
-(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         clear. hauto lq:on rew:off use:REReds.univ_inv, REReds.suc_inv. *)
-(*     + case : b => //=. *)
-(*       * qauto l:on use:T_AbsNat_Imp. *)
-(*       * qauto l:on use:T_PairNat_Imp. *)
-(*       * move => > /algo_metric_join /Sub.FromJoin. hauto l:on use:Sub.bind_nat_noconf. *)
-(*       * move => > /algo_metric_join /Sub.FromJoin. hauto lq:on use:Sub.univ_nat_noconf. *)
-(*       * hauto l:on. *)
-(*       * move /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv. *)
-(*     (* Zero *) *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on rew:off use:T_AbsZero_Imp. *)
-(*       * hauto lq: on use: T_PairZero_Imp. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.bind_inv. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.univ_inv. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.nat_inv. *)
-(*       * hauto l:on. *)
-(*       * move =>> /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.zero_inv, REReds.suc_inv. *)
-(*     (* Suc *) *)
-(*     + case : b => //=. *)
-(*       * hauto lq:on rew:off use:T_AbsSuc_Imp. *)
-(*       * hauto lq:on use:T_PairSuc_Imp. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.bind_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.nat_inv. *)
-(*       * move => > /algo_metric_join. *)
-(*         hauto lq:on rew:off use:REReds.suc_inv, REReds.zero_inv. *)
-(*       * move => a0 a1 h _ _. *)
-(*         move /algo_metric_suc : h => [j [h4 h5]]. *)
-(*         move => Γ A /Suc_Inv [h0 h1] /Suc_Inv [h2 h3]. *)
-(*         move : ih h4 h5;do!move/[apply]. *)
-(*         move => [ih _]. *)
-(*         move : ih h0 h2;do!move/[apply]. *)
-(*         move => h. apply : CE_HRed; eauto using rtc_refl. *)
-(*         by constructor. *)
-(*   - move : a b h fb fa. abstract : hnfneu. *)
-(*     move => + b. *)
-(*     case : b => //=. *)
-(*     (* NeuAbs *) *)
-(*     + move => a u halg _ nea. split => // Γ A hu /[dup] ha. *)
-(*       move /Abs_Inv => [A0][B0][_]hSu. *)
-(*       move /Sub_Bind_InvR : (hSu) => [i][A2][B2]hE. *)
-(*       have {}hu : Γ ⊢ u ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
-(*       have {}ha : Γ ⊢ PAbs a ∈ PBind PPi A2 B2 by eauto using T_Conv_E. *)
-(*       have {}hE : Γ ⊢ PBind PPi A2 B2 ∈ PUniv i *)
-(*         by hauto l:on use:regularity. *)
-(*       have {i} [j {}hE] : exists j, Γ ⊢ A2 ∈ PUniv j *)
-(*           by qauto l:on use:Bind_Inv. *)
-(*       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. *)
-(*       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. *)
-(*       apply : CE_HRed; eauto using rtc_refl. *)
-(*       eapply CE_NeuAbs => //=. *)
-(*       eapply ih; eauto. *)
-(*     (* NeuPair *) *)
-(*     + move => a0 b0 u halg _ neu. *)
-(*       split => // Γ A hu /[dup] wt. *)
-(*       move /Pair_Inv => [A0][B0][ha0][hb0]hU. *)
-(*       move /Sub_Bind_InvR : (hU) => [i][A2][B2]hE. *)
-(*       have {}wt : Γ ⊢ PPair a0 b0 ∈ PBind PSig A2 B2 by sauto lq:on. *)
-(*       have {}hu : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
-(*       move /Pair_Sig_Inv : wt => [{}ha0 {}hb0]. *)
-(*       have /T_Proj1 huL := hu. *)
-(*       have /T_Proj2 {hu} huR := hu. *)
-(*       move /algo_metric_neu_pair /(_ neu) : halg => [j [hj [hL hR]]]. *)
-(*       have heL : PProj PL u ⇔ a0 by hauto l:on. *)
-(*       eapply CE_HRed; eauto using rtc_refl. *)
-(*       apply CE_NeuPair; eauto. *)
-(*       eapply ih; eauto. *)
-(*       apply : T_Conv; eauto. *)
-(*       have {}hE : Γ ⊢ PBind PSig A2 B2 ∈ PUniv i *)
-(*         by hauto l:on use:regularity. *)
-(*       have /E_Symmetric : Γ ⊢ PProj PL u  ≡ a0 ∈ A2 by *)
-(*         hauto l:on use:coqeq_sound_mutual. *)
-(*       hauto l:on use:bind_inst. *)
-(*     (* NeuBind: Impossible *) *)
-(*     + move => b p p0 a /algo_metric_join h _ h0. exfalso. *)
-(*       move : h h0. clear. *)
-(*       move /Sub.FromJoin. *)
-(*       hauto l:on use:Sub.hne_bind_noconf. *)
-(*     (* NeuUniv: Impossible *) *)
-(*     + hauto lq:on rew:off use:DJoin.hne_univ_noconf, algo_metric_join. *)
-(*     + hauto lq:on rew:off use:DJoin.hne_nat_noconf, algo_metric_join. *)
-(*     (* Zero *) *)
-(*     + case => //=. *)
-(*       * move => i /algo_metric_join. clear. *)
-(*         hauto lq:on rew:off use:REReds.var_inv, REReds.zero_inv. *)
-(*       * move => >/algo_metric_join. clear. *)
-(*         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. *)
-(*       * move => >/algo_metric_join. clear. *)
-(*         move => h _ h2. exfalso. *)
-(*         hauto q:on use:REReds.hne_ind_inv, REReds.zero_inv. *)
-(*     (* Suc *) *)
-(*     + move => a0. *)
-(*       case => //=; move => >/algo_metric_join; clear. *)
-(*       * 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. *)
-(*       * 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 Γ 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 => 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. *)
-(*       * clear => ? ? _. exfalso. *)
-(*         hauto q:on use:REReds.var_inv, REReds.hne_ind_inv. *)
-(*     + case : b => //=; *)
-(*                    lazymatch goal with *)
-(*                    | [|- context[algo_metric _ (PApp _ _) (PApp _ _)]] => idtac *)
-(*                    | _ => move => > /algo_metric_join *)
-(*                    end. *)
-(*       * clear => *. exfalso. *)
-(*         hauto lq:on rew:off use:REReds.hne_app_inv, REReds.var_inv. *)
-(*       (* real case *) *)
-(*       * move => b1 a1 b0 a0 halg hne1 hne0 Γ A B wtA wtB. *)
-(*         move /App_Inv : wtA => [A0][B0][hb0][ha0]hS0. *)
-(*         move /App_Inv : wtB => [A1][B1][hb1][ha1]hS1. *)
-(*         move : algo_metric_app (hne0) (hne1) halg. repeat move/[apply]. *)
-(*         move => [j [hj [hal0 hal1]]]. *)
-(*         have /ih := hal0. *)
-(*         move /(_ hj). *)
-(*         move => [_ ihb]. *)
-(*         move : ihb (hne0) (hne1) hb0 hb1. repeat move/[apply]. *)
-(*         move => [hb01][C][hT0][hT1][hT2]hT3. *)
-(*         move /Sub_Bind_InvL : (hT0). *)
-(*         move => [i][A2][B2]hE. *)
-(*         have hSu20 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A0 B0 by *)
-(*           eauto using Su_Eq, Su_Transitive. *)
-(*         have hSu10 : Γ ⊢ PBind PPi A2 B2 ≲ PBind PPi A1 B1 by *)
-(*           eauto using Su_Eq, Su_Transitive. *)
-(*         have hSuA0 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
-(*         have hSuA1 : Γ ⊢ A1 ≲ A2 by sfirstorder use:Su_Pi_Proj1. *)
-(*         have ha1' : Γ ⊢ a1 ∈ A2 by eauto using T_Conv. *)
-(*         have ha0' : Γ ⊢ a0 ∈ A2 by eauto using T_Conv. *)
-(*         move : ih (hj) hal1. repeat move/[apply]. *)
-(*         move => [ih _]. *)
-(*         move : ih (ha0') (ha1'); repeat move/[apply]. *)
-(*         move => iha. *)
-(*         split. qblast. *)
-(*         exists (subst_PTm (scons a0 VarPTm) B2). *)
-(*         split. *)
-(*         apply : Su_Transitive; eauto. *)
-(*         move /E_Refl in ha0. *)
-(*         hauto l:on use:Su_Pi_Proj2. *)
-(*         have h01 : Γ ⊢ a0 ≡ a1 ∈ A2 by sfirstorder use:coqeq_sound_mutual. *)
-(*         split. *)
-(*         apply Su_Transitive with (B := subst_PTm (scons a1 VarPTm) B2). *)
-(*         move /regularity_sub0 : hSu10 => [i0]. *)
-(*         hauto l:on use:bind_inst. *)
-(*         hauto lq:on rew:off use:Su_Pi_Proj2, Su_Transitive, E_Refl. *)
-(*         split. *)
-(*         by apply : T_App; eauto using T_Conv_E. *)
-(*         apply : T_Conv; eauto. *)
-(*         apply T_App with (A := A2) (B := B2); eauto. *)
-(*         apply : T_Conv_E; eauto. *)
-(*         move /E_Symmetric in h01. *)
-(*         move /regularity_sub0 : hSu20 => [i0]. *)
-(*         sfirstorder use:bind_inst. *)
-(*       * clear => ? ? ?. exfalso. *)
-(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_proj_inv. *)
-(*       * clear => ? ? ?. exfalso. *)
-(*         hauto q:on use:REReds.hne_app_inv, REReds.hne_ind_inv. *)
-(*     + case : b => //=. *)
-(*       * move => i p h /algo_metric_join. clear => ? _ ?. exfalso. *)
-(*         hauto l:on use:REReds.hne_proj_inv, REReds.var_inv. *)
-(*       * move => > /algo_metric_join. clear => ? ? ?. exfalso. *)
-(*         hauto l:on use:REReds.hne_proj_inv, REReds.hne_app_inv. *)
-(*       (* real case *) *)
-(*       * move => p1 a1 p0 a0 /algo_metric_proj ha hne1 hne0. *)
-(*         move : ha (hne0) (hne1); repeat move/[apply]. *)
-(*         move => [? [j []]]. subst. *)
-(*         move : ih; repeat move/[apply]. *)
-(*         move => [_ ih]. *)
-(*         case : p1. *)
-(*         ** move => Γ A B ha0 ha1. *)
-(*            move /Proj1_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
-(*            move /Proj1_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
-(*            move : ih ha0 ha1 (hne0) (hne1); repeat move/[apply]. *)
-(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
-(*            split. sauto lq:on. *)
-(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
-(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            exists A2. split; eauto using Su_Sig_Proj1, Su_Transitive. *)
-(*            repeat split => //=. *)
-(*            hauto l:on use:Su_Sig_Proj1, Su_Transitive. *)
-(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
-(*            apply T_Proj1 with (B := B2); eauto using T_Conv_E. *)
-(*         ** move => Γ A B  ha0  ha1. *)
-(*            move /Proj2_Inv : ha0. move => [A0][B0][ha0]hSu0. *)
-(*            move /Proj2_Inv : ha1. move => [A1][B1][ha1]hSu1. *)
-(*            move : ih (ha0) (ha1) (hne0) (hne1); repeat move/[apply]. *)
-(*            move => [ha [C [hS0 [hS1 [wta0 wta1]]]]]. *)
-(*            split. sauto lq:on. *)
-(*            move /Sub_Bind_InvL : (hS0) => [i][A2][B2]hS2. *)
-(*            have hSu20 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A0 B0 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            have hSu21 : Γ ⊢ PBind PSig A2 B2 ≲ PBind PSig A1 B1 *)
-(*              by eauto using Su_Transitive, Su_Eq. *)
-(*            have hA20 : Γ ⊢ A2 ≲ A0 by eauto using Su_Sig_Proj1. *)
-(*            have hA21 : Γ ⊢ A2 ≲ A1 by eauto using Su_Sig_Proj1. *)
-(*            have {}wta0 : Γ ⊢ a0 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
-(*            have {}wta1 : Γ ⊢ a1 ∈ PBind PSig A2 B2 by eauto using T_Conv_E. *)
-(*            have haE : Γ ⊢ PProj PL a0 ≡ PProj PL a1 ∈ A2 *)
-(*              by sauto lq:on use:coqeq_sound_mutual. *)
-(*            exists (subst_PTm (scons (PProj PL a0) VarPTm) B2). *)
-(*            repeat split. *)
-(*            *** apply : Su_Transitive; eauto. *)
-(*                have : Γ ⊢ PProj PL a0 ≡ PProj PL a0 ∈ A2 *)
-(*                  by qauto use:regularity, E_Refl. *)
-(*                sfirstorder use:Su_Sig_Proj2. *)
-(*            *** apply : Su_Transitive; eauto. *)
-(*                sfirstorder use:Su_Sig_Proj2. *)
-(*            *** eauto using T_Proj2. *)
-(*            *** apply : T_Conv. *)
-(*                apply : T_Proj2; eauto. *)
-(*                move /E_Symmetric in haE. *)
-(*                move /regularity_sub0 in hSu21. *)
-(*                sfirstorder use:bind_inst. *)
-(*       * move => > /algo_metric_join; clear => ? ? ?. exfalso. *)
-(*         hauto q:on use:REReds.hne_proj_inv, REReds.hne_ind_inv. *)
-(*     (* ind ind case *) *)
-(*     + move => P a0 b0 c0. *)
-(*       case : b => //=; *)
-(*                    lazymatch goal with *)
-(*                    | [|- context[algo_metric _ (PInd _ _ _ _) (PInd _ _ _ _)]] => idtac *)
-(*                    | _ => move => > /algo_metric_join; clear => *; exfalso *)
-(*                    end. *)
-(*       * 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. *)
-(*       * 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. *)
-(*         move /Ind_Inv : hL => [iP0][wtP0][wta0][wtb0][wtc0]hSu0. *)
-(*         move /Ind_Inv : hR => [iP1][wtP1][wta1][wtb1][wtc1]hSu1. *)
-(*         have {}iha : a0 ∼ a1 by qauto l:on. *)
-(*         have []  : iP0 <= max iP0 iP1 /\ iP1 <= max iP0 iP1 by lia. *)
-(*         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 Δ := 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. *)
-(*         have haE : Γ ⊢ a0 ≡ a1 ∈ PNat *)
-(*           by hauto l:on use:coqeq_sound_mutual. *)
-(*         have wtΓ : ⊢ Γ by hauto l:on use:wff_mutual. *)
-(*         have hE : Γ ⊢ subst_PTm (scons PZero VarPTm) P ≡ subst_PTm (scons PZero VarPTm) P1 ∈ subst_PTm (scons PZero VarPTm) (PUniv (Nat.max iP0 iP1)). *)
-(*         eapply morphing; eauto. apply morphing_ext. by apply morphing_id. by apply T_Zero. *)
-(*         have {}wtb1 : Γ ⊢ b1 ∈ subst_PTm (scons PZero VarPTm) P *)
-(*           by eauto using T_Conv_E. *)
-(*         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 : (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. *)
-(*         eapply morphing. apply : Su_Eq. apply E_Symmetric. by eauto. *)
-(*         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. 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. *)
-(*         have hEInd : Γ ⊢ PInd P a0 b0 c0 ≡ PInd P1 a1 b1 c1 ∈ subst_PTm (scons a0 VarPTm) P by hfcrush use:coqeq_sound_mutual. *)
-(*         exists (subst_PTm (scons a0 VarPTm) P). *)
-(*         repeat split => //=; eauto with wt. *)
-(*         apply : Su_Transitive. *)
-(*         apply : Su_Sig_Proj2; eauto. apply : Su_Sig; eauto using T_Nat' with wt. *)
-(*         apply : Su_Eq. apply E_Refl. by apply T_Nat'. *)
-(*         apply : Su_Eq. apply hPE. by eauto. *)
-(*         move : hEInd. clear. hauto l:on use:regularity. *)
-(* Qed. *)
+Qed.
 
 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 Γ (a b A : PTm) : *)
-(*   Γ ⊢ a ≡ b ∈ A -> a ⇔ b. *)
-(* Proof. *)
-(*   move => h. *)
-(*   suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity. *)
-(*   eapply fundamental_theorem in h. *)
-(*   move /logrel.SemEq_SN_Join : h. *)
-(*   move => h. *)
-(*   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. *)
-(*   move => [va][vb][hva][hvb][nva][nvb]hj. *)
-(*   move /relations.rtc_nsteps : hva => [i hva]. *)
-(*   move /relations.rtc_nsteps : hvb => [j hvb]. *)
-(*   exists (i + j + size_PTm va + size_PTm vb). *)
-(*   hauto lq:on solve+:lia. *)
-(* Qed. *)
+Lemma coqeq_complete Γ (a b A : PTm) :
+  Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
+Proof.
+  move => h.
+  suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity.
+  eapply fundamental_theorem in h.
+  move /logrel.SemEq_SN_Join : h.
+  move => h.
+  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.
+  move => [va][vb][hva][hvb][nva][nvb]hj.
+  move /relations.rtc_nsteps : hva => [i hva].
+  move /relations.rtc_nsteps : hvb => [j hvb].
+  exists (i + j + size_PTm va + size_PTm vb).
+  hauto lq:on solve+:lia.
+Qed.
 
 
 Reserved Notation "a ≪ b" (at level 70).

From 96ad0a474067aae6b05c92b0ef70acc00371640b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 11 Mar 2025 00:14:43 -0400
Subject: [PATCH 197/210] Add stub for the new coqleq_complete'

---
 theories/algorithmic.v        | 456 ++++++++++++++--------------------
 theories/common.v             |   2 +
 theories/executable.v         |   8 +-
 theories/executable_correct.v |  11 -
 theories/termination.v        |  16 --
 5 files changed, 193 insertions(+), 300 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index a870f9f..d4aa29d 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -1647,26 +1647,31 @@ 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 sn_term_metric (a b : PTm) : SN a -> SN b -> exists k, term_metric k a b.
+Proof.
+  move /LoReds.FromSN => [va [ha0 ha1]].
+  move /LoReds.FromSN => [vb [hb0 hb1]].
+  eapply relations.rtc_nsteps in ha0.
+  eapply relations.rtc_nsteps in hb0.
+  hauto lq:on unfold:term_metric solve+:lia.
+Qed.
+
+Lemma sn_algo_dom a b : SN a -> SN b -> algo_dom_r a b.
+Proof.
+  move : sn_term_metric; repeat move/[apply].
+  move => [k]+.
+  eauto using term_metric_algo_dom.
+Qed.
+
 Lemma coqeq_complete Γ (a b A : PTm) :
   Γ ⊢ a ≡ b ∈ A -> a ⇔ b.
 Proof.
   move => h.
-  suff : exists k, algo_metric k a b by hauto lq:on use:coqeq_complete', regularity.
-  eapply fundamental_theorem in h.
-  move /logrel.SemEq_SN_Join : h.
-  move => h.
-  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.
-  move => [va][vb][hva][hvb][nva][nvb]hj.
-  move /relations.rtc_nsteps : hva => [i hva].
-  move /relations.rtc_nsteps : hvb => [j hvb].
-  exists (i + j + size_PTm va + size_PTm vb).
-  hauto lq:on solve+:lia.
+  have : algo_dom_r a b /\ DJoin.R a b by
+    hauto lq:on use:fundamental_theorem, logrel.SemEq_SemWt, logrel.SemWt_SN, sn_algo_dom.
+  hauto lq:on use:regularity, coqeq_complete'.
 Qed.
 
-
 Reserved Notation "a ≪ b" (at level 70).
 Reserved Notation "a ⋖ b" (at level 70).
 Inductive CoqLEq : PTm -> PTm -> Prop :=
@@ -1708,269 +1713,176 @@ Scheme coqleq_ind := Induction for CoqLEq Sort Prop
 
 Combined Scheme coqleq_mutual from coqleq_ind, coqleq_r_ind.
 
-(* 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 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 => 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.
+    apply Su_Transitive with (B := PBind PPi A1 B0).
+    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 => 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.
+    apply Su_Transitive with (B := PBind PSig A0 B1).
+    apply : Su_Sig; eauto using E_Refl, Su_Eq.
+    apply : ihB; by eauto using ctx_eq_subst_one.
+    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 => 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_algo_metric k (a b : PTm) : *)
-(*   ishne a \/ ishne b -> *)
-(*   salgo_metric k a b -> *)
-(*   algo_metric k a b. *)
-(* Proof. *)
-(*   move => h. *)
-(*   move => [i][j][va][vb][hva][hvb][nva][nvb][hS]sz. *)
-(*   rewrite/ESub.R in hS. *)
-(*   move : hS => [va'][vb'][h0][h1]h2. *)
-(*   suff : va' = vb' by sauto lq:on. *)
-(*   have {}hva : rtc RERed.R a va by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds. *)
-(*   have {}hvb : rtc RERed.R b vb by hauto lq:on use:@relations.rtc_nsteps, REReds.FromRReds, LoReds.ToRReds. *)
-(*   apply REReds.FromEReds in h0, h1. *)
-(*   have : ishne va' \/ ishne vb' by *)
-(*     hauto lq:on rew:off use:@relations.rtc_transitive, REReds.hne_preservation. *)
-(*   hauto lq:on use:Sub1.hne_refl. *)
-(* Qed. *)
+Lemma CLE_HRedL (a a' b : PTm )  :
+  HRed.R a a' ->
+  a' ≪ b ->
+  a ≪ b.
+Proof.
+  hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
+Qed.
 
-(* 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 => 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. *)
-(*     apply Su_Transitive with (B := PBind PPi A1 B0). *)
-(*     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 => 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. *)
-(*     apply Su_Transitive with (B := PBind PSig A0 B1). *)
-(*     apply : Su_Sig; eauto using E_Refl, Su_Eq. *)
-(*     apply : ihB; by eauto using ctx_eq_subst_one. *)
-(*     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 => 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 CLE_HRedR (a a' b : PTm)  :
+  HRed.R a a' ->
+  b ≪ a' ->
+  b ≪ a.
+Proof.
+  hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
+Qed.
 
-(* 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. *)
-(*   move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6. *)
-(*   case : a h0 => //=; try firstorder. *)
-(*   - inversion h0 as [|A B C D E F]; subst. *)
-(*     hauto qb:on use:ne_hne. *)
-(*     inversion E; subst => /=. *)
-(*     + hauto lq:on use:HRed.AppAbs unfold:algo_metric solve+:lia. *)
-(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*     + sfirstorder qb:on use:ne_hne. *)
-(*   - inversion h0 as [|A B C D E F]; subst. *)
-(*     hauto qb:on use:ne_hne. *)
-(*     inversion E; subst => /=. *)
-(*     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia. *)
-(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*   - inversion h0 as [|A B C D E F]; subst. *)
-(*     hauto qb:on use:ne_hne. *)
-(*     inversion E; subst => /=. *)
-(*     + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia. *)
-(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*     + sfirstorder use:ne_hne. *)
-(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*     + sfirstorder use:ne_hne. *)
-(*     + sfirstorder use:ne_hne. *)
-(* Qed. *)
-
-(* Lemma CLE_HRedL (a a' b : PTm )  : *)
-(*   HRed.R a a' -> *)
-(*   a' ≪ b -> *)
-(*   a ≪ b. *)
-(* Proof. *)
-(*   hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R. *)
-(* Qed. *)
-
-(* Lemma CLE_HRedR (a a' b : PTm)  : *)
-(*   HRed.R a a' -> *)
-(*   b ≪ a' -> *)
-(*   b ≪ a. *)
-(* Proof. *)
-(*   hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R. *)
-(* Qed. *)
+Lemma coqleq_complete' :
+  (forall a b, salgo_dom a b -> Sub.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⋖ b) /\
+  (forall a b, salgo_dom_r a b -> Sub.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ≪ b).
+Proof.
+  apply salgo_dom_mutual.
+  - move => i j /Sub.univ_inj.
+    hauto lq:on ctrs:CoqLEq.
+  - admit.
+  - admit.
+  - sfirstorder.
+  - admit.
+  - admit.
+  - hauto l:on.
+  - move => a a' b hr ha iha hj Γ A wta wtb.
+    apply : CLE_HRedL; eauto.
+    apply : iha; eauto; last by sfirstorder use:HRed.preservation.
+    apply : Sub.transitive; eauto.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply /Sub.FromJoin /DJoin.FromRRed1. by apply HRed.ToRRed.
+  - move => a b b' nfa hr h ih j Γ A wta wtb.
+    apply : CLE_HRedR; eauto.
+    apply : ih; eauto; last by eauto using HRed.preservation.
+    apply : Sub.transitive; eauto.
+    hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
+    apply /Sub.FromJoin /DJoin.FromRRed0. by apply HRed.ToRRed.
+Admitted.
 
 
-(* 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. *)
-(*   move=>[i][j][va][vb][h0] [h1][h2][h3][[v [h4 h5]]] h6. *)
-(*   case : b h1 => //=; try by firstorder. *)
-(*   - inversion 1 as [|A B C D E F]; subst. *)
-(*     hauto qb:on use:ne_hne. *)
-(*     inversion E; subst => /=. *)
-(*     + hauto q:on use:HRed.AppAbs unfold:salgo_metric solve+:lia. *)
-(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:salgo_metric solve+:lia. *)
-(*     + sfirstorder qb:on use:ne_hne. *)
-(*   - inversion 1 as [|A B C D E F]; subst. *)
-(*     hauto qb:on use:ne_hne. *)
-(*     inversion E; subst => /=. *)
-(*     + hauto lq:on use:HRed.ProjPair unfold:algo_metric solve+:lia. *)
-(*     + hauto q:on ctrs:HRed.R use: hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*   - inversion 1 as [|A B C D E F]; subst. *)
-(*     hauto qb:on use:ne_hne. *)
-(*     inversion E; subst => /=. *)
-(*     + hauto lq:on use:HRed.IndZero unfold:algo_metric solve+:lia. *)
-(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*     + sfirstorder use:ne_hne. *)
-(*     + hauto lq:on ctrs:HRed.R use:hf_hred_lored unfold:algo_metric solve+:lia. *)
-(*     + sfirstorder use:ne_hne. *)
-(*     + sfirstorder use:ne_hne. *)
-(* Qed. *)
-
-(* Lemma salgo_metric_sub k (a b : PTm) : *)
-(*   salgo_metric k a b -> *)
-(*   Sub.R a b. *)
-(* Proof. *)
-(*   rewrite /algo_metric. *)
-(*   move => [i][j][va][vb][h0][h1][h2][h3][[va' [vb' [hva [hvb hS]]]]]h5. *)
-(*   have {}h0 : rtc LoRed.R a va by hauto lq:on use:@relations.rtc_nsteps. *)
-(*   have {}h1 : rtc LoRed.R b vb by hauto lq:on use:@relations.rtc_nsteps. *)
-(*   apply REReds.FromEReds in hva,hvb. *)
-(*   apply LoReds.ToRReds in h0,h1. *)
-(*   apply REReds.FromRReds in h0,h1. *)
-(*   rewrite /Sub.R. exists va', vb'. sfirstorder use:@relations.rtc_transitive. *)
-(* Qed. *)
-
-(* 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. *)
-(*   move => [i][j][va][vb][h0][h1][h2][h3][h4]h5. *)
-(*   move : lored_nsteps_bind_inv h0 => /[apply]. *)
-(*   move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst. *)
-(*   move : lored_nsteps_bind_inv h1 => /[apply]. *)
-(*   move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst. *)
-(*   move /ESub.pi_inj : h4 => [? ?]. *)
-(*   hauto qb:on solve+:lia. *)
-(* Qed. *)
-
-(* 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. *)
-(*   move => [i][j][va][vb][h0][h1][h2][h3][h4]h5. *)
-(*   move : lored_nsteps_bind_inv h0 => /[apply]. *)
-(*   move => [i0][i1][a2][b2][?][?][?][ha02]hb02. subst. *)
-(*   move : lored_nsteps_bind_inv h1 => /[apply]. *)
-(*   move => [j0][j1][a3][b3][?][?][?][ha13]hb13. subst. *)
-(*   move /ESub.sig_inj : h4 => [? ?]. *)
-(*   hauto qb:on solve+:lia. *)
-(* Qed. *)
-
-(* 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 a b. *)
-(*   elim /Wf_nat.lt_wf_ind. *)
-(*   move => n ih. *)
-(*   move => a b /[dup] h /salgo_metric_case. *)
-(*   (* Cases where a and b can take steps *) *)
-(*   case; cycle 1. *)
-(*   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. *)
-(*   (* Cases where neither a nor b can take steps *) *)
-(*   case => fb; case => fa. *)
-(*   - case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
-(*     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
-(*       * move => p0 A0 B0 _ p1 A1 B1 _. *)
-(*         move => h. *)
-(*         have ? : p1 = p0 by *)
-(*           hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj. *)
-(*         subst. *)
-(*         case : p0 h => //=. *)
-(*         ** move /salgo_metric_pi. *)
-(*            move => [j [hj [hA hB]]] Γ i. *)
-(*            move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0]. *)
-(*            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 : (cons A0 Γ) ⊢ B1 ∈ PUniv i *)
-(*              by hauto l:on. *)
-(*            eauto using ctx_eq_subst_one. *)
-(*         ** move /salgo_metric_sig. *)
-(*            move => [j [hj [hA hB]]] Γ i. *)
-(*            move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0]. *)
-(*            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 : (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. *)
-(*       * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf. *)
-(*       * move => _ > _ /salgo_metric_sub. *)
-(*         hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv inv:Sub1.R. *)
-(*       * hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R. *)
-(*     + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
-(*       * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf. *)
-(*       * move => *. econstructor; eauto using rtc_refl. *)
-(*         hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong. *)
-(*       * hauto lq:on rew:off use:REReds.univ_inv, REReds.nat_inv, salgo_metric_sub inv:Sub1.R. *)
-(*       * hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R. *)
-(*       * hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R. *)
-(*   (* Both cases are impossible *) *)
-(*     + case : b fb => //=. *)
-(*       * hauto lq:on use:T_AbsNat_Imp. *)
-(*       * hauto lq:on use:T_PairNat_Imp. *)
-(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.bind_inv, salgo_metric_sub inv:Sub1.R. *)
-(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R. *)
-(*       * hauto l:on. *)
-(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R. *)
-(*       * hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R. *)
-(*     + move => ? ? Γ i /Zero_Inv. *)
-(*       clear. *)
-(*       move /synsub_to_usub => [_ [_ ]]. *)
-(*       hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R. *)
-(*     + move => ? _ _ Γ i /Suc_Inv => [[_]]. *)
-(*       move /synsub_to_usub. *)
-(*       hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R. *)
-(*   - have {}h : DJoin.R a b by *)
-(*        hauto lq:on use:salgo_metric_algo_metric, algo_metric_join. *)
-(*     case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
-(*     + hauto lq:on use:DJoin.hne_bind_noconf. *)
-(*     + hauto lq:on use:DJoin.hne_univ_noconf. *)
-(*     + hauto lq:on use:DJoin.hne_nat_noconf. *)
-(*     + move => _ _ Γ i _. *)
-(*       move /Zero_Inv. *)
-(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
-(*     + move => p _ _ Γ i _ /Suc_Inv. *)
-(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
-(*   - have {}h : DJoin.R b a by *)
-(*       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric. *)
-(*     case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'. *)
-(*     + hauto lq:on use:DJoin.hne_bind_noconf. *)
-(*     + hauto lq:on use:DJoin.hne_univ_noconf. *)
-(*     + hauto lq:on use:DJoin.hne_nat_noconf. *)
-(*     + move => _ _ Γ i. *)
-(*       move /Zero_Inv. *)
-(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
-(*     + move => p _ _ Γ i /Suc_Inv. *)
-(*       hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R. *)
-(*   - move => Γ i ha hb. *)
-(*     econstructor; eauto using rtc_refl. *)
-(*     apply CLE_NeuNeu. move {ih}. *)
-(*     have {}h : algo_metric n a b by *)
-(*       hauto lq:on use:salgo_metric_algo_metric. *)
-(*     eapply coqeq_complete'; eauto. *)
-(* Qed. *)
+  move : k a b.
+  elim /Wf_nat.lt_wf_ind.
+  move => n ih.
+  move => a b /[dup] h /salgo_metric_case.
+  (* Cases where a and b can take steps *)
+  case; cycle 1.
+  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.
+  (* Cases where neither a nor b can take steps *)
+  case => fb; case => fa.
+  - case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+      * move => p0 A0 B0 _ p1 A1 B1 _.
+        move => h.
+        have ? : p1 = p0 by
+          hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj.
+        subst.
+        case : p0 h => //=.
+        ** move /salgo_metric_pi.
+           move => [j [hj [hA hB]]] Γ i.
+           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
+           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 : (cons A0 Γ) ⊢ B1 ∈ PUniv i
+             by hauto l:on.
+           eauto using ctx_eq_subst_one.
+        ** move /salgo_metric_sig.
+           move => [j [hj [hA hB]]] Γ i.
+           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
+           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 : (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.
+      * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf.
+      * move => _ > _ /salgo_metric_sub.
+        hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
+    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+      * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
+      * move => *. econstructor; eauto using rtc_refl.
+        hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong.
+      * hauto lq:on rew:off use:REReds.univ_inv, REReds.nat_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
+  (* Both cases are impossible *)
+    + case : b fb => //=.
+      * hauto lq:on use:T_AbsNat_Imp.
+      * hauto lq:on use:T_PairNat_Imp.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.bind_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto l:on.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
+      * hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
+    + move => ? ? Γ i /Zero_Inv.
+      clear.
+      move /synsub_to_usub => [_ [_ ]].
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
+    + move => ? _ _ Γ i /Suc_Inv => [[_]].
+      move /synsub_to_usub.
+      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
+  - have {}h : DJoin.R a b by
+       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join.
+    case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+    + hauto lq:on use:DJoin.hne_bind_noconf.
+    + hauto lq:on use:DJoin.hne_univ_noconf.
+    + hauto lq:on use:DJoin.hne_nat_noconf.
+    + move => _ _ Γ i _.
+      move /Zero_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
+    + move => p _ _ Γ i _ /Suc_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
+  - have {}h : DJoin.R b a by
+      hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric.
+    case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
+    + hauto lq:on use:DJoin.hne_bind_noconf.
+    + hauto lq:on use:DJoin.hne_univ_noconf.
+    + hauto lq:on use:DJoin.hne_nat_noconf.
+    + move => _ _ Γ i.
+      move /Zero_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
+    + move => p _ _ Γ i /Suc_Inv.
+      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
+  - move => Γ i ha hb.
+    econstructor; eauto using rtc_refl.
+    apply CLE_NeuNeu. move {ih}.
+    have {}h : algo_metric n a b by
+      hauto lq:on use:salgo_metric_algo_metric.
+    eapply coqeq_complete'; eauto.
+Qed.
 
 (* Lemma coqleq_complete Γ (a b : PTm) : *)
 (*   Γ ⊢ a ≲ b -> a ≪ b. *)
diff --git a/theories/common.v b/theories/common.v
index 39713ba..35267fc 100644
--- a/theories/common.v
+++ b/theories/common.v
@@ -403,6 +403,8 @@ with salgo_dom_r : PTm -> PTm -> Prop :=
 Scheme salgo_ind := Induction for salgo_dom Sort Prop
   with salgor_ind := Induction for salgo_dom_r Sort Prop.
 
+Combined Scheme salgo_dom_mutual from salgo_ind, salgor_ind.
+
 Lemma hf_no_hred (a b : PTm) :
   ishf a ->
   HRed.R a b ->
diff --git a/theories/executable.v b/theories/executable.v
index e62e3c3..558aa75 100644
--- a/theories/executable.v
+++ b/theories/executable.v
@@ -341,4 +341,10 @@ Proof.
     sfirstorder use:hne_no_hred, hf_no_hred.
 Qed.
 
-#[export]Hint Rewrite check_sub_hredl check_sub_hredr check_sub_nfnf check_sub_univ_univ check_sub_pi_pi check_sub_sig_sig : ce_prop.
+Lemma check_sub_neuneu a b i a0 : check_sub a b (S_NeuNeu a b i a0) = check_equal a b a0.
+Proof. destruct a,b => //=. Qed.
+
+Lemma check_sub_conf a b n n0 i : check_sub a b (S_Conf a b n n0 i) = false.
+Proof. destruct a,b=>//=; ecrush inv:BTag. Qed.
+
+#[export]Hint Rewrite check_sub_neuneu check_sub_conf check_sub_hredl check_sub_hredr check_sub_nfnf check_sub_univ_univ check_sub_pi_pi check_sub_sig_sig : ce_prop.
diff --git a/theories/executable_correct.v b/theories/executable_correct.v
index 4b76638..40debce 100644
--- a/theories/executable_correct.v
+++ b/theories/executable_correct.v
@@ -180,17 +180,6 @@ Qed.
 
 Ltac simp_sub := with_strategy opaque [check_equal] simpl in *.
 
-Combined Scheme salgo_dom_mutual from salgo_ind, salgor_ind.
-
-(* Reusing algo_dom results in an inefficient proof, but i'll brute force it so i can get the result more quickly *)
-Lemma check_sub_neuneu a b i a0 : check_sub a b (S_NeuNeu a b i a0) = check_equal a b a0.
-Proof. destruct a,b => //=. Qed.
-
-Lemma check_sub_conf a b n n0 i : check_sub a b (S_Conf a b n n0 i) = false.
-Proof. destruct a,b=>//=; ecrush inv:BTag. Qed.
-
-Hint Rewrite check_sub_neuneu check_sub_conf : ce_prop.
-
 Lemma check_sub_sound :
   (forall a b (h : salgo_dom a b), check_sub a b h -> a ⋖ b) /\
     (forall a b (h : salgo_dom_r a b), check_sub_r a b h -> a ≪ b).
diff --git a/theories/termination.v b/theories/termination.v
index 41400ef..c5bc37e 100644
--- a/theories/termination.v
+++ b/theories/termination.v
@@ -3,19 +3,3 @@ From Hammer Require Import Tactics.
 Require Import ssreflect ssrbool.
 From stdpp Require Import relations (nsteps (..), rtc(..)).
 Import Psatz.
-
-Lemma sn_term_metric (a b : PTm) : SN a -> SN b -> exists k, term_metric k a b.
-Proof.
-  move /LoReds.FromSN => [va [ha0 ha1]].
-  move /LoReds.FromSN => [vb [hb0 hb1]].
-  eapply relations.rtc_nsteps in ha0.
-  eapply relations.rtc_nsteps in hb0.
-  hauto lq:on unfold:term_metric solve+:lia.
-Qed.
-
-Lemma sn_algo_dom a b : SN a -> SN b -> algo_dom_r a b.
-Proof.
-  move : sn_term_metric; repeat move/[apply].
-  move => [k]+.
-  eauto using term_metric_algo_dom.
-Qed.

From 8dbef3e29ed6f0a6c31fb26f0b2934d6b7fd9f5a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 11 Mar 2025 16:24:57 -0400
Subject: [PATCH 198/210] Finish refactoring

---
 theories/algorithmic.v | 243 +++++++++++++++++++----------------------
 1 file changed, 112 insertions(+), 131 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index d4aa29d..d9711e1 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -947,6 +947,18 @@ Proof.
   hauto lq:on use:synsub_to_usub, Sub.bind_univ_noconf, Abs_Inv.
 Qed.
 
+Lemma T_ZeroUniv_Imp' Γ i :
+  Γ ⊢ PZero ∈ PUniv i -> False.
+Proof.
+  hauto lq:on use:synsub_to_usub, Sub.univ_nat_noconf, Zero_Inv.
+Qed.
+
+Lemma T_SucUniv_Imp' Γ (a : PTm) i :
+  Γ ⊢ PSuc a ∈ PUniv i -> False.
+Proof.
+  hauto lq:on use:synsub_to_usub, Sub.univ_nat_noconf, Suc_Inv.
+Qed.
+
 Lemma T_PairUniv_Imp' Γ (a b : PTm) i  :
   Γ ⊢ PPair a b ∈ PUniv i -> False.
 Proof.
@@ -1760,18 +1772,95 @@ Proof.
   hauto lq:on ctrs:rtc, CoqLEq_R inv:CoqLEq_R.
 Qed.
 
+Lemma subvar_inj (i j : nat) :
+    Sub.R (VarPTm i) (VarPTm j) -> i = j.
+Proof.
+  rewrite /Sub.R.
+  move => [c][d][h0][h1]h2.
+  apply REReds.var_inv in h0, h1. subst.
+  inversion h2; by subst.
+Qed.
+
+Lemma algo_dom_hf_hne (a b : PTm) :
+  algo_dom a b ->
+  (ishf a \/ ishne a) /\ (ishf b \/ ishne b).
+Proof.
+  inversion 1;subst => //=; by sfirstorder b:on.
+Qed.
+
+Lemma algo_dom_neu_neu_nonconf a b :
+  algo_dom a b ->
+  neuneu_nonconf a b ->
+  ishne a /\ ishne b.
+Proof.
+  move /algo_dom_hf_hne => h.
+  move => h1.
+  destruct a,b => //=; sfirstorder b:on.
+Qed.
+
 Lemma coqleq_complete' :
-  (forall a b, salgo_dom a b -> Sub.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ⋖ b) /\
-  (forall a b, salgo_dom_r a b -> Sub.R a b -> forall Γ A, Γ ⊢ a ∈ A -> Γ ⊢ b ∈ A -> a ≪ b).
+  (forall a b, salgo_dom a b -> Sub.R a b -> forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> a ⋖ b) /\
+  (forall a b, salgo_dom_r a b -> Sub.R a b -> forall Γ i, Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv i -> a ≪ b).
 Proof.
   apply salgo_dom_mutual.
   - move => i j /Sub.univ_inj.
     hauto lq:on ctrs:CoqLEq.
-  - admit.
-  - admit.
+  - move => A0 A1 B0 B1 hA ihA hB ihB /Sub.bind_inj. move => [_][hjA]hjB Γ i.
+    move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
+    have {}ihA : A1 ≪ A0 by hauto l:on.
+    constructor => //.
+    have ihA' : Γ ⊢ A1 ≲ A0 by hauto l:on use:coqleq_sound_mutual.
+    suff : (cons A1 Γ) ⊢ B0 ∈ PUniv i
+      by hauto l:on.
+    eauto using ctx_eq_subst_one.
+  - move => A0 A1 B0 B1 hA ihA hB ihB /Sub.bind_inj. move => [_][hjA]hjB Γ i.
+    move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
+    have {}ihA : A0 ≪ A1 by hauto l:on.
+    constructor => //.
+    have ihA' : Γ ⊢ A0 ≲ A1 by hauto l:on use:coqleq_sound_mutual.
+    suff : (cons A0 Γ) ⊢ B1 ∈ PUniv i
+      by hauto l:on.
+    eauto using ctx_eq_subst_one.
   - sfirstorder.
-  - admit.
-  - admit.
+  - move => a b hconf hdom.
+    have [? ?] : ishne a /\ ishne b by sfirstorder use:algo_dom_neu_neu_nonconf.
+    move => h. apply Sub.ToJoin in h; last by tauto.
+    move => Γ i ha hb.
+    apply CLE_NeuNeu. hauto q:on use:coqeq_complete'.
+  - move => [:neunf] a b.
+    case => ha; case => hb.
+    move : ha hb.
+    + case : a => //=; try solve [intros; exfalso; eauto using T_AbsUniv_Imp', T_PairUniv_Imp', T_ZeroUniv_Imp', T_SucUniv_Imp'].
+      * case : b => //=; try solve [intros; exfalso; eauto using T_AbsUniv_Imp', T_PairUniv_Imp', T_ZeroUniv_Imp', T_SucUniv_Imp'].
+        case => + + []//=.
+        hauto lq:on rew:off use:Sub.bind_inj.
+        hauto lq:on rew:off use:Sub.bind_inj.
+        hauto lq:on use:Sub.bind_univ_noconf.
+        hauto lq:on use:Sub.nat_bind_noconf.
+      * case : b => //=; try solve [intros; exfalso; eauto using T_AbsUniv_Imp', T_PairUniv_Imp', T_ZeroUniv_Imp', T_SucUniv_Imp'].
+        hauto lq:on use:Sub.univ_bind_noconf.
+        hauto lq:on use:Sub.nat_univ_noconf.
+      * case : b => //=; try solve [intros; exfalso; eauto using T_AbsUniv_Imp', T_PairUniv_Imp', T_ZeroUniv_Imp', T_SucUniv_Imp'].
+        hauto lq:on use:Sub.bind_nat_noconf.
+        hauto lq:on use:Sub.univ_nat_noconf.
+    + move => h0 h1.
+      apply Sub.ToJoin in h1; last by tauto.
+      move => Γ i wta wtb. exfalso.
+      abstract : neunf a b ha hb h0 h1 Γ i wta wtb.
+      case : a ha h0 h1 wta => //=; eauto using T_AbsUniv_Imp', T_PairUniv_Imp', T_ZeroUniv_Imp', T_SucUniv_Imp'.
+      sfirstorder use: DJoin.hne_bind_noconf.
+      sfirstorder use: DJoin.hne_univ_noconf.
+      sfirstorder use:DJoin.hne_nat_noconf.
+    + move => h0 h1. apply Sub.ToJoin in h1; last by tauto.
+      hauto drew:off use:DJoin.symmetric, stm_conf_sym.
+    + move => h0 h1 Γ i wta wtb.
+      apply CLE_NeuNeu.
+      apply Sub.ToJoin in h1; last by tauto.
+      eapply coqeq_complete'; eauto.
+      apply algo_dom_r_algo_dom.
+      sfirstorder use:hne_no_hred.
+      sfirstorder use:hne_no_hred.
+      hauto lq:on use:sn_algo_dom, logrel.SemWt_SN, fundamental_theorem.
   - hauto l:on.
   - move => a a' b hr ha iha hj Γ A wta wtb.
     apply : CLE_HRedL; eauto.
@@ -1785,130 +1874,22 @@ Proof.
     apply : Sub.transitive; eauto.
     hauto lq:on use:fundamental_theorem, logrel.SemWt_SN.
     apply /Sub.FromJoin /DJoin.FromRRed0. by apply HRed.ToRRed.
-Admitted.
-
-
-  move : k a b.
-  elim /Wf_nat.lt_wf_ind.
-  move => n ih.
-  move => a b /[dup] h /salgo_metric_case.
-  (* Cases where a and b can take steps *)
-  case; cycle 1.
-  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.
-  (* Cases where neither a nor b can take steps *)
-  case => fb; case => fa.
-  - case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-      * move => p0 A0 B0 _ p1 A1 B1 _.
-        move => h.
-        have ? : p1 = p0 by
-          hauto lq:on rew:off use:salgo_metric_sub, Sub.bind_inj.
-        subst.
-        case : p0 h => //=.
-        ** move /salgo_metric_pi.
-           move => [j [hj [hA hB]]] Γ i.
-           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
-           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 : (cons A0 Γ) ⊢ B1 ∈ PUniv i
-             by hauto l:on.
-           eauto using ctx_eq_subst_one.
-        ** move /salgo_metric_sig.
-           move => [j [hj [hA hB]]] Γ i.
-           move /Bind_Univ_Inv => [hA1 hB1] /Bind_Univ_Inv [hA0 hB0].
-           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 : (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.
-      * hauto lq:on use:salgo_metric_sub, Sub.nat_bind_noconf.
-      * move => _ > _ /salgo_metric_sub.
-        hauto lq:on rew:off use:REReds.bind_inv, REReds.zero_inv inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.bind_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
-    + case : b fb => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-      * hauto lq:on use:salgo_metric_sub, Sub.univ_bind_noconf.
-      * move => *. econstructor; eauto using rtc_refl.
-        hauto lq:on use:salgo_metric_sub, Sub.univ_inj, CLE_UnivCong.
-      * hauto lq:on rew:off use:REReds.univ_inv, REReds.nat_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.univ_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.suc_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
-  (* Both cases are impossible *)
-    + case : b fb => //=.
-      * hauto lq:on use:T_AbsNat_Imp.
-      * hauto lq:on use:T_PairNat_Imp.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.bind_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto l:on.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.zero_inv, salgo_metric_sub inv:Sub1.R.
-      * hauto lq:on rew:off use:REReds.nat_inv, REReds.suc_inv, salgo_metric_sub inv:Sub1.R.
-    + move => ? ? Γ i /Zero_Inv.
-      clear.
-      move /synsub_to_usub => [_ [_ ]].
-      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
-    + move => ? _ _ Γ i /Suc_Inv => [[_]].
-      move /synsub_to_usub.
-      hauto lq:on rew:off use:REReds.nat_inv, REReds.univ_inv inv:Sub1.R.
-  - have {}h : DJoin.R a b by
-       hauto lq:on use:salgo_metric_algo_metric, algo_metric_join.
-    case : b fb h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-    + hauto lq:on use:DJoin.hne_bind_noconf.
-    + hauto lq:on use:DJoin.hne_univ_noconf.
-    + hauto lq:on use:DJoin.hne_nat_noconf.
-    + move => _ _ Γ i _.
-      move /Zero_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-    + move => p _ _ Γ i _ /Suc_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-  - have {}h : DJoin.R b a by
-      hauto lq:on use:salgo_metric_algo_metric, algo_metric_join, DJoin.symmetric.
-    case : a fa h => //=; try hauto depth:1 lq:on use:T_AbsUniv_Imp', T_PairUniv_Imp'.
-    + hauto lq:on use:DJoin.hne_bind_noconf.
-    + hauto lq:on use:DJoin.hne_univ_noconf.
-    + hauto lq:on use:DJoin.hne_nat_noconf.
-    + move => _ _ Γ i.
-      move /Zero_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-    + move => p _ _ Γ i /Suc_Inv.
-      hauto q:on use:REReds.nat_inv, REReds.univ_inv, synsub_to_usub inv:Sub1.R.
-  - move => Γ i ha hb.
-    econstructor; eauto using rtc_refl.
-    apply CLE_NeuNeu. move {ih}.
-    have {}h : algo_metric n a b by
-      hauto lq:on use:salgo_metric_algo_metric.
-    eapply coqeq_complete'; eauto.
+Qed.
+Lemma coqleq_complete Γ (a b : PTm) :
+  Γ ⊢ a ≲ b -> a ≪ b.
+Proof.
+  move => h.
+  have : salgo_dom_r a b /\ Sub.R a b by
+    hauto lq:on use:fundamental_theorem, logrel.SemLEq_SemWt, logrel.SemWt_SN, sn_algo_dom, algo_dom_salgo_dom.
+  hauto lq:on use:regularity, coqleq_complete'.
 Qed.
 
-(* Lemma coqleq_complete Γ (a b : PTm) : *)
-(*   Γ ⊢ a ≲ b -> a ≪ b. *)
-(* Proof. *)
-(*   move => h. *)
-(*   suff : exists k, salgo_metric k a b by hauto lq:on use:coqleq_complete', regularity. *)
-(*   eapply fundamental_theorem in h. *)
-(*   move /logrel.SemLEq_SN_Sub : h. *)
-(*   move => h. *)
-(*   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. *)
-(*   move => [va][vb][hva][hvb][nva][nvb]hj. *)
-(*   move /relations.rtc_nsteps : hva => [i hva]. *)
-(*   move /relations.rtc_nsteps : hvb => [j hvb]. *)
-(*   exists (i + j + size_PTm va + size_PTm vb). *)
-(*   hauto lq:on solve+:lia. *)
-(* Qed. *)
-
-(* Lemma coqleq_sound : forall Γ (a b : PTm) i j, *)
-(*     Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv j -> a ≪ b -> Γ ⊢ a ≲ b. *)
-(* Proof. *)
-(*   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. *)
-(*   sfirstorder use:coqleq_sound_mutual. *)
-(* Qed. *)
+Lemma coqleq_sound : forall Γ (a b : PTm) i j,
+    Γ ⊢ a ∈ PUniv i -> Γ ⊢ b ∈ PUniv j -> a ≪ b -> Γ ⊢ a ≲ b.
+Proof.
+  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.
+  sfirstorder use:coqleq_sound_mutual.
+Qed.

From f9d3a620f4c6d920465f3b84bd2f3441c7cb07bc Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Tue, 11 Mar 2025 16:27:31 -0400
Subject: [PATCH 199/210] Remove itauto dependency as it introduces weird
 axioms

---
 theories/algorithmic.v |  5 ++---
 theories/fp_red.v      | 14 +++++++-------
 theories/logrel.v      | 16 ++++++++--------
 3 files changed, 17 insertions(+), 18 deletions(-)

diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index d9711e1..184872d 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -5,7 +5,6 @@ Require Import ssreflect ssrbool.
 Require Import Psatz.
 From stdpp Require Import relations (rtc(..), nsteps(..)).
 Require Import Coq.Logic.FunctionalExtensionality.
-Require Import Cdcl.Itauto.
 
 Module HRed.
   Lemma ToRRed (a b : PTm) : HRed.R a b -> RRed.R a b.
@@ -98,7 +97,7 @@ Proof.
   - move => u v hC /andP [h0 h1] /synsub_to_usub ?.
     exfalso.
     suff : SNe (PApp u v) by hauto l:on use:Sub.bind_sne_noconf.
-    eapply ne_nf_embed => //=. itauto.
+    eapply ne_nf_embed => //=. sfirstorder b:on.
   - move => p0 p1 hC  h  ?. exfalso.
     have {hC} : Γ ⊢ PPair p0 p1 ∈ PUniv i by hauto l:on use:regularity.
     hauto lq:on use:Sub.bind_univ_noconf, synsub_to_usub, Pair_Inv.
@@ -199,7 +198,7 @@ Proof.
   - move => u v hC /andP [h0 h1] /synsub_to_usub ?.
     exfalso.
     suff : SNe (PApp u v) by hauto l:on use:Sub.sne_bind_noconf.
-    eapply ne_nf_embed => //=. itauto.
+    eapply ne_nf_embed => //=. sfirstorder b:on.
   - move => p0 p1 hC  h  ?. exfalso.
     have {hC} : Γ ⊢ PPair p0 p1 ∈ PUniv i by hauto l:on use:regularity.
     hauto lq:on use:Sub.bind_univ_noconf, synsub_to_usub, Pair_Inv.
diff --git a/theories/fp_red.v b/theories/fp_red.v
index 9d8718b..382f25b 100644
--- a/theories/fp_red.v
+++ b/theories/fp_red.v
@@ -10,7 +10,7 @@ From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn).
 From Hammer Require Import Tactics.
 Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common.
 Require Import Btauto.
-Require Import Cdcl.Itauto.
+
 
 Ltac2 spec_refl () :=
   List.iter
@@ -2575,7 +2575,7 @@ Module LoReds.
     ~~ ishf a.
   Proof.
     move : hf_preservation. repeat move/[apply].
-    case : a; case : b => //=; itauto.
+    case : a; case : b => //=; sfirstorder b:on.
   Qed.
 
   #[local]Ltac solve_s_rec :=
@@ -2633,7 +2633,7 @@ Module LoReds.
     rtc LoRed.R (PSuc a0) (PSuc a1).
   Proof. solve_s. Qed.
 
-  Local Ltac triv := simpl in *; itauto.
+  Local Ltac triv := simpl in *; sfirstorder b:on.
 
   Lemma FromSN_mutual :
     (forall (a : PTm) (_ : SNe a), exists v, rtc LoRed.R a v /\ ne v) /\
@@ -3048,7 +3048,7 @@ Module DJoin.
     have {h0 h1 h2 h3} : isbind c /\ isuniv c by
       hauto l:on use:REReds.bind_preservation,
           REReds.univ_preservation.
-    case : c => //=; itauto.
+    case : c => //=; sfirstorder b:on.
   Qed.
 
   Lemma hne_univ_noconf (a b : PTm) :
@@ -3078,7 +3078,7 @@ Module DJoin.
   Proof.
     move => [c [h0 h1]] h2 h3.
     have : ishne c /\ isnat c by sfirstorder use:REReds.hne_preservation, REReds.nat_preservation.
-    clear. case : c => //=; itauto.
+    clear. case : c => //=; sfirstorder b:on.
   Qed.
 
   Lemma bind_inj p0 p1 (A0 A1 : PTm) B0 B1 :
@@ -3594,7 +3594,7 @@ Module Sub.
       hauto l:on use:REReds.bind_preservation,
             REReds.univ_preservation, Sub1.bind_preservation, Sub1.univ_preservation.
     move : h2 h3. clear.
-    case : c => //=; itauto.
+    case : c => //=; sfirstorder b:on.
   Qed.
 
   Lemma univ_bind_noconf (a b : PTm) :
@@ -3605,7 +3605,7 @@ Module Sub.
       hauto l:on use:REReds.bind_preservation,
             REReds.univ_preservation, Sub1.bind_preservation, Sub1.univ_preservation.
     move : h2 h3. clear.
-    case : c => //=; itauto.
+    case : c => //=; sfirstorder b:on.
   Qed.
 
   Lemma bind_inj p0 p1 (A0 A1 : PTm) B0 B1 :
diff --git a/theories/logrel.v b/theories/logrel.v
index a479f81..d4b4bd9 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -6,7 +6,7 @@ Require Import ssreflect ssrbool.
 Require Import Logic.PropExtensionality (propositional_extensionality).
 From stdpp Require Import relations (rtc(..), rtc_subrel).
 Import Psatz.
-Require Import Cdcl.Itauto.
+
 
 Definition ProdSpace (PA : PTm -> Prop)
   (PF : PTm -> (PTm -> Prop) -> Prop) b : Prop :=
@@ -355,7 +355,7 @@ Proof.
       move => [H [/DJoin.FromRedSNs h [h1 h0]]].
       have {h}{}hAB : Sub.R A H by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have {}h0 : SNe H.
-      suff : ~ isbind H /\ ~ isuniv H /\ ~ isnat H by itauto.
+      suff : ~ isbind H /\ ~ isuniv H /\ ~ isnat H by sfirstorder b:on.
       move : hA hAB. clear. hauto lq:on db:noconf.
       move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
       tauto.
@@ -365,7 +365,7 @@ Proof.
       move => [H [/DJoin.FromRedSNs h [h1 h0]]].
       have {h}{}hAB : Sub.R H A by qauto l:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have {}h0 : SNe H.
-      suff : ~ isbind H /\ ~ isuniv H /\ ~ isnat H by itauto.
+      suff : ~ isbind H /\ ~ isuniv H /\ ~ isnat H by sfirstorder b:on.
       move : hAB hA h0. clear. hauto lq:on db:noconf.
       move : InterpUniv_SNe_inv h1 h0. repeat move/[apply]. move => ?. subst.
       tauto.
@@ -375,7 +375,7 @@ Proof.
       have {hU} {}h : Sub.R (PBind p A B) H
         by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have{h0} : isbind H.
-      suff : ~ SNe H /\ ~ isuniv H /\ ~ isnat H by itauto.
+      suff : ~ SNe H /\ ~ isuniv H /\ ~ isnat H by sfirstorder b:on.
       have : isbind (PBind p A B) by scongruence.
       move : h. clear. hauto l:on db:noconf.
       case : H h1 h => //=.
@@ -394,7 +394,7 @@ Proof.
       have {hU} {}h : Sub.R H (PBind p A B)
         by move : hU hU' h; clear; hauto q:on use:Sub.FromJoin, DJoin.symmetric, Sub.transitive.
       have{h0} : isbind H.
-      suff : ~ SNe H /\ ~ isuniv H /\ ~ isnat H by itauto.
+      suff : ~ SNe H /\ ~ isuniv H /\ ~ isnat H by sfirstorder b:on.
       have : isbind (PBind p A B) by scongruence.
       move : h. clear. move : (PBind p A B). hauto lq:on db:noconf.
       case : H h1 h => //=.
@@ -413,7 +413,7 @@ Proof.
       move => [H [/DJoin.FromRedSNs h [h1 h0]]].
       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.
+      suff : ~ isbind H /\ ~ isuniv H /\ ~ SNe H by sfirstorder b:on.
       have : @isnat PNat by simpl.
       move : h0 hAB. clear. qauto l:on db:noconf.
       case : H h1 hAB h0 => //=.
@@ -424,7 +424,7 @@ Proof.
       move => [H [/DJoin.FromRedSNs h [h1 h0]]].
       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.
+      suff : ~ isbind H /\ ~ isuniv H /\ ~ SNe H by sfirstorder b:on.
       have : @isnat PNat by simpl.
       move : h0 hAB. clear. qauto l:on db:noconf.
       case : H h1 hAB h0 => //=.
@@ -1058,7 +1058,7 @@ Definition Γ_sub' Γ Δ := forall ρ, ρ_ok Δ ρ -> ρ_ok Γ ρ.
 Definition Γ_eq' Γ Δ := forall ρ, ρ_ok Δ ρ <-> ρ_ok Γ ρ.
 
 Lemma Γ_sub'_refl Γ : Γ_sub' Γ Γ.
-  rewrite /Γ_sub'. itauto. Qed.
+  rewrite /Γ_sub'. sfirstorder b:on. Qed.
 
 Lemma Γ_sub'_cons Γ Δ A B i j :
   Sub.R B A ->

From 4b2bbeea6fe7c456c440e00672e8ccbdb161483b Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Wed, 16 Apr 2025 23:57:00 -0400
Subject: [PATCH 200/210] Add SE_FunExt

---
 theories/logrel.v | 14 ++++++++++++++
 1 file changed, 14 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index d4b4bd9..d78d459 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1516,6 +1516,20 @@ Proof.
   rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
 Qed.
 
+Lemma SE_FunExt Γ (a b : PTm) A B i :
+  Γ ⊨ PBind PPi A B ∈ PUniv i ->
+  Γ ⊨ a ∈ PBind PPi A B ->
+  Γ ⊨ b ∈ PBind PPi A B ->
+  A :: Γ ⊨ PApp (ren_PTm shift a) (VarPTm var_zero) ≡ PApp (ren_PTm shift b) (VarPTm var_zero) ∈ B ->
+  Γ ⊨ a ≡ b ∈ PBind PPi A B.
+Proof.
+  move => hpi ha hb he.
+  move : SE_Abs (hpi) he. repeat move/[apply]. move => he.
+  have /SE_Symmetric : Γ ⊨ PAbs (PApp (ren_PTm shift a) (VarPTm var_zero)) ≡ a ∈ PBind PPi A B by eauto using SE_AppEta.
+  have : Γ ⊨ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B by eauto using SE_AppEta.
+  eauto using SE_Transitive.
+Qed.
+
 Lemma SE_Nat Γ (a b : PTm) :
   Γ ⊨ a ≡ b ∈ PNat ->
   Γ ⊨ PSuc a ≡ PSuc b ∈ PNat.

From ef8feb63c3c9df872f769e86c676c3aa04fcdaca Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 17 Apr 2025 15:07:14 -0400
Subject: [PATCH 201/210] Redefine semantic context well-formedness as an
 inductive

---
 theories/logrel.v | 53 ++++++++++++++++++++++++++---------------------
 1 file changed, 29 insertions(+), 24 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index d78d459..f85b32e 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -619,10 +619,6 @@ Proof.
   split; apply : InterpUniv_cumulative; eauto.
 Qed.
 
-(* Semantic context wellformedness *)
-Definition SemWff Γ := forall (i : nat) A, lookup i Γ A -> exists j, Γ ⊨ A ∈ PUniv j.
-Notation "⊨ Γ" := (SemWff Γ) (at level 70).
-
 Lemma ρ_ok_id Γ :
   ρ_ok Γ VarPTm.
 Proof.
@@ -763,6 +759,34 @@ Proof.
   move : weakening_Sem h0 h1; repeat move/[apply]. done.
 Qed.
 
+Reserved Notation "⊨ Γ"  (at level 70).
+
+Inductive SemWff : list PTm -> Prop :=
+| SemWff_nil :
+  ⊨ nil
+| SemWff_cons Γ A i :
+  ⊨ Γ ->
+  Γ ⊨ A ∈ PUniv i ->
+  (* -------------- *)
+  ⊨ (cons A Γ)
+where "⊨ Γ" := (SemWff Γ).
+
+(* Semantic context wellformedness *)
+Lemma SemWff_lookup Γ :
+  ⊨ Γ ->
+  forall (i : nat) A, lookup i Γ A -> exists j, Γ ⊨ A ∈ PUniv j.
+Proof.
+  move => h. elim : Γ / h.
+  - by inversion 1.
+  - move => Γ A i hΓ ihΓ hA  j B.
+    elim /lookup_inv => //=_.
+    + move => ? ? ? [*]. subst.
+      eauto using weakening_Sem_Univ.
+    + move => i0 Γ0 A0 B0 hl ? [*]. subst.
+      move : ihΓ hl => /[apply]. move => [j hA0].
+      eauto using weakening_Sem_Univ.
+Qed.
+
 Lemma SemWt_SN Γ (a : PTm) A :
   Γ ⊨ a ∈ A ->
   SN a /\ SN A.
@@ -784,25 +808,6 @@ Lemma SemLEq_SN_Sub Γ (a b : PTm) :
   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 nil.
-Proof. hfcrush inv:lookup. Qed.
-
-Lemma SemWff_cons Γ (A : PTm) i :
-    ⊨ Γ ->
-    Γ ⊨ A ∈ PUniv i ->
-    (* -------------- *)
-    ⊨ (cons A Γ).
-Proof.
-  move => h h0.
-  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' Γ (i : nat) A j :
   lookup i Γ A ->
@@ -819,7 +824,7 @@ Lemma ST_Var Γ (i : nat) A :
   ⊨ Γ ->
   lookup i Γ A ->
   Γ ⊨ VarPTm i ∈ A.
-Proof. hauto l:on use:ST_Var' unfold:SemWff. Qed.
+Proof. hauto l:on use:ST_Var', SemWff_lookup. Qed.
 
 Lemma InterpUniv_Bind_nopf p i (A : PTm) B PA :
   ⟦ A ⟧ i ↘ PA ->

From a367591e9a67e5cc33cfd6e4d85a9d932fa74f6a Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 17 Apr 2025 15:15:58 -0400
Subject: [PATCH 202/210] Add extensional version of pair equality rules

---
 theories/logrel.v | 16 ++++++++++++++++
 1 file changed, 16 insertions(+)

diff --git a/theories/logrel.v b/theories/logrel.v
index f85b32e..fa50b9c 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1521,6 +1521,22 @@ Proof.
   rewrite /DJoin.R. hauto lq:on ctrs:rtc,RERed.R.
 Qed.
 
+Lemma SE_PairExt Γ (a b : PTm) A B i :
+  Γ ⊨ PBind PSig A B ∈ PUniv i ->
+  Γ ⊨ a ∈ PBind PSig A B ->
+  Γ ⊨ b ∈ PBind PSig A B ->
+  Γ ⊨ PProj PL a ≡ PProj PL b ∈ A ->
+  Γ ⊨ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B ->
+  Γ ⊨ a ≡ b ∈ PBind PSig A B.
+Proof.
+  move => h0 ha hb h1 h2.
+  suff h : Γ ⊨ a ≡ PPair (PProj PL a) (PProj PR a) ∈ PBind PSig A B /\
+           Γ ⊨ PPair (PProj PL b) (PProj PR b) ≡ b ∈ PBind PSig A B /\
+           Γ ⊨ PPair (PProj PL a) (PProj PR a) ≡ PPair (PProj PL b) (PProj PR b) ∈ PBind PSig A B
+    by decompose [and] h; eauto using SE_Transitive, SE_Symmetric.
+  eauto 20 using SE_PairEta, SE_Symmetric, SE_Pair.
+Qed.
+
 Lemma SE_FunExt Γ (a b : PTm) A B i :
   Γ ⊨ PBind PPi A B ∈ PUniv i ->
   Γ ⊨ a ∈ PBind PPi A B ->

From e1fc6b609dd69dd4ae69f6b1146eab095a923e69 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 17 Apr 2025 15:22:45 -0400
Subject: [PATCH 203/210] Add the extensional representation of pair&abs
 equality rules

---
 theories/logrel.v |  2 +-
 theories/typing.v | 36 +++++++++++-------------------------
 2 files changed, 12 insertions(+), 26 deletions(-)

diff --git a/theories/logrel.v b/theories/logrel.v
index fa50b9c..e11659e 100644
--- a/theories/logrel.v
+++ b/theories/logrel.v
@@ -1701,4 +1701,4 @@ Qed.
 
 #[export]Hint Resolve ST_Var ST_Bind ST_Abs ST_App ST_Pair ST_Proj1 ST_Proj2 ST_Univ ST_Conv
   SE_Refl SE_Symmetric SE_Transitive SE_Bind SE_Abs SE_App SE_Proj1 SE_Proj2
-  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong SE_IndZero SE_IndSuc SE_SucCong : sem.
+  SE_Conv SSu_Pi_Proj1 SSu_Pi_Proj2 SSu_Sig_Proj1 SSu_Sig_Proj2 SSu_Eq SSu_Transitive SSu_Pi SSu_Sig SemWff_nil SemWff_cons SSu_Univ SE_AppAbs SE_ProjPair1 SE_ProjPair2 SE_AppEta SE_PairEta ST_Nat ST_Ind ST_Suc ST_Zero SE_IndCong SE_SucCong SE_IndZero SE_IndSuc SE_SucCong SE_PairExt SE_FunExt : sem.
diff --git a/theories/typing.v b/theories/typing.v
index ae78747..e911d51 100644
--- a/theories/typing.v
+++ b/theories/typing.v
@@ -89,23 +89,12 @@ with Eq : list PTm -> PTm -> PTm -> PTm -> Prop :=
   (cons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
   Γ ⊢ PBind p A0 B0 ≡ PBind p A1 B1 ∈ PUniv i
 
-| E_Abs Γ (a b : PTm) A B i :
-  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
-  (cons A Γ) ⊢ a ≡ b ∈ B ->
-  Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B
-
 | E_App Γ i (b0 b1 a0 a1 : PTm) A B :
   Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
   Γ ⊢ b0 ≡ b1 ∈ PBind PPi A B ->
   Γ ⊢ a0 ≡ a1 ∈ A ->
   Γ ⊢ PApp b0 a0 ≡ PApp b1 a1 ∈ subst_PTm (scons a0 VarPTm) B
 
-| E_Pair Γ (a0 a1 b0 b1 : PTm) A B i :
-  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
-  Γ ⊢ a0 ≡ a1 ∈ A ->
-  Γ ⊢ b0 ≡ b1 ∈ subst_PTm (scons a0 VarPTm) B ->
-  Γ ⊢ PPair a0 b0 ≡ PPair a1 b1 ∈ PBind PSig A B
-
 | E_Proj1 Γ (a b : PTm) A B :
   Γ ⊢ a ≡ b ∈ PBind PSig A B ->
   Γ ⊢ PProj PL a ≡ PProj PL b ∈ A
@@ -164,16 +153,20 @@ with Eq : list PTm -> PTm -> PTm -> PTm -> Prop :=
   (cons P (cons PNat Γ)) ⊢ c ∈ ren_PTm shift (subst_PTm (scons (PSuc (VarPTm var_zero)) (funcomp VarPTm shift) ) P) ->
   Γ ⊢ PInd P (PSuc a) b c ≡ (subst_PTm (scons (PInd P a b c) (scons a VarPTm)) c) ∈ subst_PTm (scons (PSuc a) VarPTm) P
 
-(* Eta *)
-| E_AppEta Γ (b : PTm) A B i :
-  Γ ⊢ PBind PPi A B ∈ (PUniv i) ->
+| E_FunExt Γ (a b : PTm) A B i :
+  Γ ⊢ PBind PPi A B ∈ PUniv i ->
+  Γ ⊢ a ∈ PBind PPi A B ->
   Γ ⊢ b ∈ PBind PPi A B ->
-  Γ ⊢ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B
+  A :: Γ ⊢ PApp (ren_PTm shift a) (VarPTm var_zero) ≡ PApp (ren_PTm shift b) (VarPTm var_zero) ∈ B ->
+  Γ ⊢ a ≡ b ∈ PBind PPi A B
 
-| E_PairEta Γ (a : PTm ) A B i :
-  Γ ⊢ PBind PSig A B ∈ (PUniv i) ->
+| E_PairExt Γ (a b : 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
+  Γ ⊢ b ∈ PBind PSig A B ->
+  Γ ⊢ PProj PL a ≡ PProj PL b ∈ A ->
+  Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B ->
+  Γ ⊢ a ≡ b ∈ PBind PSig A B
 
 with LEq : list PTm -> PTm -> PTm -> Prop :=
 (* Structural *)
@@ -242,10 +235,3 @@ Scheme wf_ind := Induction for Wff Sort Prop
   with le_ind := Induction for LEq Sort Prop.
 
 Combined Scheme wt_mutual from wf_ind, wt_ind, eq_ind, le_ind.
-
-(* Lemma lem : *)
-(*   (forall n (Γ : fin n -> PTm n), ⊢ Γ -> ...) /\ *)
-(*   (forall n Γ (a A : PTm n), Γ ⊢ a ∈ A -> ...)  /\ *)
-(*   (forall n Γ (a b A : PTm n), Γ ⊢ a ≡ b ∈ A -> ...) /\ *)
-(*   (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> ...). *)
-(* Proof. apply wt_mutual. ... *)

From 7c985fa58e1080e94dcbd877b9bd64df22c7183e Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Thu, 17 Apr 2025 16:42:57 -0400
Subject: [PATCH 204/210] Minor

---
 theories/structural.v | 28 ++++++++++------------------
 1 file changed, 10 insertions(+), 18 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index ccb4c6f..76f679b 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -137,12 +137,12 @@ Lemma E_ProjPair2' Γ (a b : PTm) A B i U :
   Γ ⊢ PProj PR (PPair a b) ≡ b ∈ U.
 Proof. move => ->. apply E_ProjPair2. Qed.
 
-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 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' Γ (a0 a1 A0 A1 : PTm ) B0 B1 U T :
   U = subst_PTm (scons a0 VarPTm) B0 ->
@@ -222,17 +222,7 @@ 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 => Γ 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 => Γ 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 => Γ P0 P1 a0 a1 b0 b1 c0 c1 i hP0 ihP0 hP ihP ha iha hb ihb hc ihc.
     move => Δ ξ hΔ hξ [:hP' hren].
@@ -302,8 +292,10 @@ Proof.
     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.
+  - move => Γ a b A B i hPi ihPi ha iha hb ihb he0 ihe1 Δ ξ hΔ hξ.
+    apply : E_FunExt; eauto. move : ihe1. asimpl. apply.
+    admit.
+    by apply renaming_up.
   - qauto l:on use:E_PairEta.
   - hauto lq:on ctrs:LEq.
   - qauto l:on ctrs:LEq.

From 87b84794b470eecb4b056a8e1c0d4fb49834645c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 18 Apr 2025 14:13:46 -0400
Subject: [PATCH 205/210] Fix structural rules

---
 theories/structural.v | 37 +++++++++++--------------------------
 1 file changed, 11 insertions(+), 26 deletions(-)

diff --git a/theories/structural.v b/theories/structural.v
index 76f679b..207447d 100644
--- a/theories/structural.v
+++ b/theories/structural.v
@@ -294,9 +294,10 @@ Proof.
     by asimpl.
   - move => Γ a b A B i hPi ihPi ha iha hb ihb he0 ihe1 Δ ξ hΔ hξ.
     apply : E_FunExt; eauto. move : ihe1. asimpl. apply.
-    admit.
+    hfcrush use:Bind_Inv.
     by apply renaming_up.
-  - qauto l:on use:E_PairEta.
+  - move => Γ a b A B i hPi ihPi ha iha hb ihb hL ihL hR ihR Δ ξ hΔ hξ.
+    apply : E_PairExt; eauto. move : ihR. asimpl. by apply.
   - hauto lq:on ctrs:LEq.
   - qauto l:on ctrs:LEq.
   - hauto lq:on ctrs:Wff use:renaming_up, Su_Pi.
@@ -439,17 +440,7 @@ Proof.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on ctrs:Eq.
   - hauto lq:on rew:off use:E_Bind', Wff_Cons, morphing_up.
-  - 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 => Γ 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 => Γ P0 P1 a0 a1 b0 b1 c0 c1 i hP0 ihP0 hP ihP ha iha hb ihb hc ihc.
@@ -516,9 +507,14 @@ Proof.
     move /(_ Ξ  (up_PTm_PTm (up_PTm_PTm ξ)) hΞ) : ihc.
     move /(_ ltac:(sauto l:on use:morphing_up)).
     asimpl. substify. by asimpl.
-  - move => *.
-    apply : E_AppEta'; eauto. by asimpl.
-  - qauto l:on use:E_PairEta.
+  - move => Γ a b A B i hPi ihPi ha iha hb ihb he0 ihe1 Δ ξ hΔ hξ.
+    move  : ihPi (hΔ) (hξ); repeat move/[apply]. move /[dup] /Bind_Inv => ihPi ?.
+    decompose [and ex] ihPi.
+    apply : E_FunExt; eauto. move : ihe1. asimpl. apply.
+    by eauto with wt.
+    by eauto using morphing_up with wt.
+  - move => Γ a b A B i hPi ihPi ha iha hb ihb hL ihL hR ihR Δ ξ hΔ hξ.
+    apply : E_PairExt; eauto. move : ihR. asimpl. by apply.
   - hauto lq:on ctrs:LEq.
   - qauto l:on ctrs:LEq.
   - hauto lq:on ctrs:Wff use:morphing_up, Su_Pi.
@@ -659,7 +655,6 @@ Proof.
     sfirstorder.
     apply : ctx_eq_subst_one; eauto using Su_Eq, E_Symmetric.
     hauto lq:on.
-  - hauto lq:on ctrs:Wt,Eq.
   - move => n i b0 b1 a0 a1 A B hP _ hb [ihb0 [ihb1 [i0 ihb2]]]
              ha [iha0 [iha1 [i1 iha2]]].
     repeat split.
@@ -669,7 +664,6 @@ Proof.
     move /E_Symmetric in ha.
     by eauto using bind_inst.
     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 => Γ i a b A B hS _ hab [iha][ihb][j]ihs.
     repeat split => //=; eauto with wt.
@@ -718,15 +712,6 @@ Proof.
     subst Ξ.
     move : morphing_wt hc; repeat move/[apply]. asimpl. by apply.
     sauto lq:on use:substing_wt.
-  - move => Γ b A B i hP _ hb [i0 ihb].
-    repeat split => //=; eauto with wt.
-    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. asimpl. by rewrite subst_scons_id.
-    apply T_Var. sfirstorder use:wff_mutual.
-    apply here.
-  - hauto lq:on ctrs:Wt.
   - move => Γ A B C hA [i [ihA0 ihA1]] hC [j [ihC0 ihC1]].
     have ? : ⊢ Γ by sfirstorder use:wff_mutual.
     exists (max i j).

From d9282fb815a8b2c18df40b573c5c2f969e908b07 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 18 Apr 2025 15:42:40 -0400
Subject: [PATCH 206/210] Finish preservation

---
 theories/preservation.v | 88 ++++++++++++++++++++++++++++++++++++++++-
 1 file changed, 86 insertions(+), 2 deletions(-)

diff --git a/theories/preservation.v b/theories/preservation.v
index c8a2106..2794909 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -1,4 +1,4 @@
-Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common typing structural fp_red.
+Require Import Autosubst2.core Autosubst2.unscoped Autosubst2.syntax common typing structural fp_red admissible.
 From Hammer Require Import Tactics.
 Require Import ssreflect.
 Require Import Psatz.
@@ -112,6 +112,45 @@ Proof.
   eauto using T_Conv.
 Qed.
 
+Lemma T_Eta Γ A a B :
+  A :: Γ ⊢ a ∈ B ->
+  A :: Γ ⊢ PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a)) (VarPTm var_zero) ∈ B.
+Proof.
+  move => ha.
+  have hΓ' : ⊢ A :: Γ by sfirstorder use:wff_mutual.
+  have [i hA] : exists i, Γ ⊢ A ∈ PUniv i by hauto lq:on inv:Wff.
+  have hΓ : ⊢ Γ by hauto lq:on inv:Wff.
+  eapply T_App' with (B  := ren_PTm (upRen_PTm_PTm shift) B). by asimpl; rewrite subst_scons_id.
+  apply T_Abs. eapply renaming; eauto; cycle 1. apply renaming_up. apply renaming_shift.
+  econstructor; eauto. sauto l:on use:renaming.
+  apply T_Var => //.
+  by constructor.
+Qed.
+
+Lemma E_Abs Γ a b A B :
+  A :: Γ ⊢ a ≡ b ∈ B ->
+  Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B.
+Proof.
+  move => h.
+  have [i hA] : exists i, Γ ⊢ A ∈ PUniv i by hauto l:on use:wff_mutual inv:Wff.
+  have [j hB] : exists j, A :: Γ ⊢ B ∈ PUniv j by hauto l:on use:regularity.
+  have hΓ : ⊢ Γ by sfirstorder use:wff_mutual.
+  have hΓ' : ⊢ A::Γ by eauto with wt.
+  set k := max i j.
+  have [? ?] : i <= k /\ j <= k by lia.
+  have {}hA : Γ ⊢ A ∈ PUniv k by hauto l:on use:T_Conv, Su_Univ.
+  have {}hB : A :: Γ ⊢ B ∈ PUniv k by hauto lq:on use:T_Conv, Su_Univ.
+  have hPi : Γ ⊢ PBind PPi A B ∈ PUniv k by eauto with wt.
+  apply : E_FunExt; eauto with wt.
+  hauto lq:on rew:off use:regularity, T_Abs.
+  hauto lq:on rew:off use:regularity, T_Abs.
+  apply : E_Transitive => /=. apply E_AppAbs.
+  hauto lq:on use:T_Eta, regularity.
+  apply /E_Symmetric /E_Transitive. apply E_AppAbs.
+  hauto lq:on use:T_Eta, regularity.
+  asimpl. rewrite !subst_scons_id. by apply E_Symmetric.
+Qed.
+
 Lemma E_ProjPair1 : forall (a b : PTm) Γ (A : PTm),
     Γ ⊢ PProj PL (PPair a b) ∈ A -> Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A.
 Proof.
@@ -125,6 +164,51 @@ Proof.
   apply : E_ProjPair1; eauto.
 Qed.
 
+Lemma E_ProjPair2 : forall (a b : PTm) Γ (A : PTm),
+    Γ ⊢ PProj PR (PPair a b) ∈ A -> Γ ⊢ PProj PR (PPair a b) ≡ b ∈ A.
+Proof.
+  move => a b Γ A. move /Proj2_Inv.
+  move => [A0][B0][ha]h.
+  have hr := ha.
+  move /Pair_Inv : ha => [A1][B1][ha][hb]hU.
+  have [i hSig] : exists i, Γ ⊢ PBind PSig A1 B1 ∈ PUniv i by sfirstorder use:regularity.
+  have /E_Symmetric : Γ ⊢ (PProj PL (PPair a b)) ≡ a ∈ A1 by eauto using  E_ProjPair1 with wt.
+  move /Su_Sig_Proj2 : hU. repeat move/[apply]. move => hB.
+  apply : E_Conv; eauto.
+  apply : E_Conv; eauto.
+  apply : E_ProjPair2; eauto.
+Qed.
+
+Lemma E_Pair Γ a0 b0 a1 b1 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.
+Proof.
+  move => hSig ha hb.
+  have [ha0 ha1] : Γ ⊢ a0 ∈ A /\ Γ ⊢ a1 ∈ A by hauto l:on use:regularity.
+  have [hb0 hb1] : Γ ⊢ b0 ∈ subst_PTm (scons a0 VarPTm) B /\
+                     Γ ⊢ b1 ∈ subst_PTm (scons a0 VarPTm) B by hauto l:on use:regularity.
+  have hp : Γ ⊢ PPair a0 b0 ∈ PBind PSig A B by eauto with wt.
+  have hp' : Γ ⊢ PPair a1 b1 ∈ PBind PSig A B. econstructor; eauto with wt; [idtac].
+  apply : T_Conv; eauto. apply : Su_Sig_Proj2; by eauto using Su_Eq, E_Refl.
+  have ea : Γ ⊢ PProj PL (PPair a0 b0) ≡ a0 ∈ A by eauto with wt.
+  have : Γ ⊢ PProj PR (PPair a0 b0) ≡ b0 ∈ subst_PTm (scons a0 VarPTm) B by eauto with wt.
+  have : Γ ⊢ PProj PL (PPair a1 b1) ≡ a1 ∈ A by eauto using E_ProjPair1 with wt.
+  have : Γ ⊢ PProj PR (PPair a1 b1) ≡ b1 ∈ subst_PTm (scons a0 VarPTm) B.
+  apply : E_Conv; eauto using E_ProjPair2 with wt.
+  apply : Su_Sig_Proj2. apply /Su_Eq /E_Refl. eassumption.
+  apply : E_Transitive. apply E_ProjPair1. by eauto with wt.
+  by eauto using E_Symmetric.
+  move => *.
+  apply : E_PairExt; eauto using E_Symmetric, E_Transitive.
+  apply : E_Conv. by eauto using E_Symmetric, E_Transitive.
+  apply : Su_Sig_Proj2. apply /Su_Eq /E_Refl. eassumption.
+  apply : E_Transitive. by eauto. apply /E_Symmetric /E_Transitive.
+  by eauto using E_ProjPair1.
+  eauto.
+Qed.
+
 Lemma Suc_Inv Γ (a : PTm) A :
   Γ ⊢ PSuc a ∈ A -> Γ ⊢ a ∈ PNat /\ Γ ⊢ PNat ≲ A.
 Proof.
@@ -159,7 +243,7 @@ Proof.
         apply : Su_Sig_Proj2; eauto.
       move : hA0 => /[swap]. move : Su_Transitive. repeat move/[apply].
       move {hS}.
-      move => ?. apply : E_Conv; eauto. apply : E_ProjPair2; eauto.
+      move => ?. apply : E_Conv; eauto. apply : typing.E_ProjPair2; eauto.
   - hauto lq:on use:Ind_Inv, E_Conv, E_IndZero.
   - move => P a b c Γ A.
     move /Ind_Inv.

From 43daff1b278f9bf21cdc89f9a69faca2b0e0b82c Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 18 Apr 2025 15:46:07 -0400
Subject: [PATCH 207/210] Fix soundness

---
 theories/soundness.v | 2 +-
 1 file changed, 1 insertion(+), 1 deletion(-)

diff --git a/theories/soundness.v b/theories/soundness.v
index 7f86852..877e3fb 100644
--- a/theories/soundness.v
+++ b/theories/soundness.v
@@ -7,7 +7,7 @@ Theorem fundamental_theorem :
   (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].
+  apply wt_mutual; eauto with sem.
   Unshelve. all : exact 0.
 Qed.
 

From 2b92845e3e00c4a3f1385c4f1c30cdd353b1e319 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <liuyiyun@seas.upenn.edu>
Date: Fri, 18 Apr 2025 16:38:34 -0400
Subject: [PATCH 208/210] Add E_AppEta

---
 theories/admissible.v   | 112 ++++++++++++++++++++++++++++++++++++++++
 theories/algorithmic.v  |   3 +-
 theories/preservation.v |  60 ---------------------
 3 files changed, 114 insertions(+), 61 deletions(-)

diff --git a/theories/admissible.v b/theories/admissible.v
index 830ceec..48c7cea 100644
--- a/theories/admissible.v
+++ b/theories/admissible.v
@@ -16,6 +16,70 @@ Proof.
   hauto lq:on rew:off inv:Wff use:T_Bind', typing.T_Abs.
 Qed.
 
+
+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 : Γ u U / hu => //=.
+  - move => Γ b a A B hb _ ha _ b0 a0 [*]. subst.
+    exists A,B.
+    repeat split => //=.
+    have [i] : exists i, Γ ⊢ PBind PPi A B ∈  PUniv i by sfirstorder use:regularity.
+    hauto lq:on use:bind_inst, E_Refl.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+
+Lemma Abs_Inv Γ (a : PTm) U :
+  Γ ⊢ PAbs a ∈ U ->
+  exists A B, (cons A Γ) ⊢ a ∈ B /\ Γ ⊢ PBind PPi A B ≲ U.
+Proof.
+  move E : (PAbs a) => u hu.
+  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.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+
+
+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 => a b Γ A ha.
+  move /App_Inv : ha.
+  move => [A0][B0][ha][hb]hS.
+  move /Abs_Inv : ha => [A1][B1][ha]hS0.
+  have hb' := hb.
+  move /E_Refl in hb.
+  have hS1 : Γ ⊢ A0 ≲ A1 by sfirstorder use:Su_Pi_Proj1.
+  have [i hPi] : exists i, Γ ⊢ PBind PPi A1 B1 ∈ PUniv i by sfirstorder use:regularity_sub0.
+  move : Su_Pi_Proj2 hS0 hb; repeat move/[apply].
+  move : hS => /[swap]. move : Su_Transitive. repeat move/[apply].
+  move => h.
+  apply : E_Conv; eauto.
+  apply : E_AppAbs; eauto.
+  eauto using T_Conv.
+Qed.
+
+Lemma T_Eta Γ A a B :
+  A :: Γ ⊢ a ∈ B ->
+  A :: Γ ⊢ PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a)) (VarPTm var_zero) ∈ B.
+Proof.
+  move => ha.
+  have hΓ' : ⊢ A :: Γ by sfirstorder use:wff_mutual.
+  have [i hA] : exists i, Γ ⊢ A ∈ PUniv i by hauto lq:on inv:Wff.
+  have hΓ : ⊢ Γ by hauto lq:on inv:Wff.
+  eapply T_App' with (B  := ren_PTm (upRen_PTm_PTm shift) B). by asimpl; rewrite subst_scons_id.
+  apply T_Abs. eapply renaming; eauto; cycle 1. apply renaming_up. apply renaming_shift.
+  econstructor; eauto. sauto l:on use:renaming.
+  apply T_Var => //.
+  by constructor.
+Qed.
+
 Lemma E_Bind Γ i p (A0 A1 : PTm) B0 B1 :
   Γ ⊢ A0 ≡ A1 ∈ PUniv i ->
   (cons A0 Γ) ⊢ B0 ≡ B1 ∈ PUniv i ->
@@ -46,3 +110,51 @@ Proof.
   have [i] : exists i, Γ ⊢ PBind PSig A B ∈ PUniv i by hauto l:on use:regularity.
   move : E_Proj2 h; by repeat move/[apply].
 Qed.
+
+Lemma E_FunExt Γ (a b : PTm) A B :
+  Γ ⊢ a ∈ PBind PPi A B ->
+  Γ ⊢ b ∈ PBind PPi A B ->
+  A :: Γ ⊢ PApp (ren_PTm shift a) (VarPTm var_zero) ≡ PApp (ren_PTm shift b) (VarPTm var_zero) ∈ B ->
+  Γ ⊢ a ≡ b ∈ PBind PPi A B.
+Proof.
+  hauto l:on use:regularity, E_FunExt.
+Qed.
+
+Lemma E_AppEta Γ (b : PTm) A B :
+  Γ ⊢ b ∈ PBind PPi A B ->
+  Γ ⊢ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B.
+Proof.
+  move => h.
+  have [i hPi] : exists i, Γ ⊢ PBind PPi A B ∈ PUniv i by sfirstorder use:regularity.
+  have [j [hA hB]] : exists i, Γ ⊢ A ∈ PUniv i /\ A :: Γ ⊢ B ∈ PUniv i by hauto lq:on use:Bind_Inv.
+  have {i} {}hPi : Γ ⊢ PBind PPi A B ∈ PUniv j by sfirstorder use:T_Bind, wff_mutual.
+  have hΓ : ⊢ A :: Γ by hauto lq:on use:Bind_Inv, Wff_Cons'.
+  have hΓ' :  ⊢ ren_PTm shift A :: A :: Γ by sauto lq:on use:renaming, renaming_shift inv:Wff.
+  apply E_FunExt; eauto.
+  apply T_Abs.
+  eapply T_App' with (A := ren_PTm shift A) (B := ren_PTm (upRen_PTm_PTm shift) B). by asimpl; rewrite subst_scons_id.
+  change (PBind _ _ _) with (ren_PTm shift (PBind PPi A B)).
+
+  eapply renaming; eauto.
+  apply renaming_shift.
+  constructor => //.
+  by constructor.
+
+  apply : E_Transitive. simpl.
+  apply E_AppAbs' with (i := j)(A := ren_PTm shift A) (B := ren_PTm (upRen_PTm_PTm shift) B); eauto.
+  by asimpl; rewrite subst_scons_id.
+  hauto q:on use:renaming, renaming_shift.
+  constructor => //.
+  by constructor.
+  asimpl.
+  eapply T_App' with (A := ren_PTm shift (ren_PTm shift A)) (B := ren_PTm (upRen_PTm_PTm shift) (ren_PTm (upRen_PTm_PTm shift) B)); cycle 2.
+  constructor. econstructor; eauto. sauto lq:on use:renaming, renaming_shift.
+  by constructor. asimpl. substify. by asimpl.
+  have -> : PBind PPi (ren_PTm shift (ren_PTm shift A)) (ren_PTm (upRen_PTm_PTm shift) (ren_PTm (upRen_PTm_PTm shift) B))=  (ren_PTm (funcomp shift shift) (PBind PPi A B)) by asimpl.
+  eapply renaming; eauto. admit.
+  asimpl. renamify.
+  eapply E_App' with (A := ren_PTm shift A) (B := ren_PTm (upRen_PTm_PTm shift) B). by asimpl; rewrite subst_scons_id.
+  hauto q:on use:renaming, renaming_shift.
+  sauto lq:on use:renaming, renaming_shift, E_Refl.
+  constructor. constructor=>//. constructor.
+Admitted.
diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index 184872d..cca7cfa 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -586,7 +586,8 @@ Proof.
     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.
     eauto using E_Symmetric, Su_Eq.
-    apply : E_Abs; eauto. hauto l:on use:regularity.
+    apply : E_Abs; eauto.
+
     apply iha.
     move /Su_Pi_Proj2_Var in hp0.
     apply : T_Conv; eauto.
diff --git a/theories/preservation.v b/theories/preservation.v
index 2794909..2722ee7 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -4,32 +4,6 @@ Require Import ssreflect.
 Require Import Psatz.
 Require Import Coq.Logic.FunctionalExtensionality.
 
-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 : Γ u U / hu => //=.
-  - move => Γ b a A B hb _ ha _ b0 a0 [*]. subst.
-    exists A,B.
-    repeat split => //=.
-    have [i] : exists i, Γ ⊢ PBind PPi A B ∈  PUniv i by sfirstorder use:regularity.
-    hauto lq:on use:bind_inst, E_Refl.
-  - hauto lq:on rew:off ctrs:LEq.
-Qed.
-
-Lemma Abs_Inv Γ (a : PTm) U :
-  Γ ⊢ PAbs a ∈ U ->
-  exists A B, (cons A Γ) ⊢ a ∈ B /\ Γ ⊢ PBind PPi A B ≲ U.
-Proof.
-  move E : (PAbs a) => u hu.
-  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.
-  - hauto lq:on rew:off ctrs:LEq.
-Qed.
-
 Lemma Proj1_Inv Γ (a : PTm ) U :
   Γ ⊢ PProj PL a ∈ U ->
   exists A B, Γ ⊢ a ∈ PBind PSig A B /\ Γ ⊢ A ≲ U.
@@ -93,40 +67,6 @@ Proof.
   - hauto lq:on rew:off ctrs:LEq.
 Qed.
 
-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 => a b Γ A ha.
-  move /App_Inv : ha.
-  move => [A0][B0][ha][hb]hS.
-  move /Abs_Inv : ha => [A1][B1][ha]hS0.
-  have hb' := hb.
-  move /E_Refl in hb.
-  have hS1 : Γ ⊢ A0 ≲ A1 by sfirstorder use:Su_Pi_Proj1.
-  have [i hPi] : exists i, Γ ⊢ PBind PPi A1 B1 ∈ PUniv i by sfirstorder use:regularity_sub0.
-  move : Su_Pi_Proj2 hS0 hb; repeat move/[apply].
-  move : hS => /[swap]. move : Su_Transitive. repeat move/[apply].
-  move => h.
-  apply : E_Conv; eauto.
-  apply : E_AppAbs; eauto.
-  eauto using T_Conv.
-Qed.
-
-Lemma T_Eta Γ A a B :
-  A :: Γ ⊢ a ∈ B ->
-  A :: Γ ⊢ PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a)) (VarPTm var_zero) ∈ B.
-Proof.
-  move => ha.
-  have hΓ' : ⊢ A :: Γ by sfirstorder use:wff_mutual.
-  have [i hA] : exists i, Γ ⊢ A ∈ PUniv i by hauto lq:on inv:Wff.
-  have hΓ : ⊢ Γ by hauto lq:on inv:Wff.
-  eapply T_App' with (B  := ren_PTm (upRen_PTm_PTm shift) B). by asimpl; rewrite subst_scons_id.
-  apply T_Abs. eapply renaming; eauto; cycle 1. apply renaming_up. apply renaming_shift.
-  econstructor; eauto. sauto l:on use:renaming.
-  apply T_Var => //.
-  by constructor.
-Qed.
-
 Lemma E_Abs Γ a b A B :
   A :: Γ ⊢ a ≡ b ∈ B ->
   Γ ⊢ PAbs a ≡ PAbs b ∈ PBind PPi A B.

From 8e083aad4b83f8013a80b38f7fb6df69197241ac Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 19 Apr 2025 00:07:48 -0400
Subject: [PATCH 209/210] Add renaming_comp

---
 theories/admissible.v | 12 ++++++++++--
 1 file changed, 10 insertions(+), 2 deletions(-)

diff --git a/theories/admissible.v b/theories/admissible.v
index 48c7cea..e7f0769 100644
--- a/theories/admissible.v
+++ b/theories/admissible.v
@@ -120,6 +120,14 @@ Proof.
   hauto l:on use:regularity, E_FunExt.
 Qed.
 
+Lemma renaming_comp Γ Δ Ξ ξ0 ξ1 :
+  renaming_ok Γ Δ ξ0 -> renaming_ok Δ Ξ ξ1 ->
+  renaming_ok Γ Ξ (funcomp ξ0 ξ1).
+  rewrite /renaming_ok => h0 h1 i A.
+  move => {}/h1 {}/h0.
+  by asimpl.
+Qed.
+
 Lemma E_AppEta Γ (b : PTm) A B :
   Γ ⊢ b ∈ PBind PPi A B ->
   Γ ⊢ PAbs (PApp (ren_PTm shift b) (VarPTm var_zero)) ≡ b ∈ PBind PPi A B.
@@ -151,10 +159,10 @@ Proof.
   constructor. econstructor; eauto. sauto lq:on use:renaming, renaming_shift.
   by constructor. asimpl. substify. by asimpl.
   have -> : PBind PPi (ren_PTm shift (ren_PTm shift A)) (ren_PTm (upRen_PTm_PTm shift) (ren_PTm (upRen_PTm_PTm shift) B))=  (ren_PTm (funcomp shift shift) (PBind PPi A B)) by asimpl.
-  eapply renaming; eauto. admit.
+  eapply renaming; eauto. by eauto using renaming_shift, renaming_comp.
   asimpl. renamify.
   eapply E_App' with (A := ren_PTm shift A) (B := ren_PTm (upRen_PTm_PTm shift) B). by asimpl; rewrite subst_scons_id.
   hauto q:on use:renaming, renaming_shift.
   sauto lq:on use:renaming, renaming_shift, E_Refl.
   constructor. constructor=>//. constructor.
-Admitted.
+Qed.

From 0e04a7076b069ab311fde2239ac8940eabcbdf60 Mon Sep 17 00:00:00 2001
From: Yiyun Liu <yliu8e5f7@protonmail.com>
Date: Sat, 19 Apr 2025 00:24:09 -0400
Subject: [PATCH 210/210] Prove PairEta

---
 theories/admissible.v   | 96 +++++++++++++++++++++++++++++++++++++++++
 theories/algorithmic.v  |  6 +--
 theories/preservation.v | 73 -------------------------------
 3 files changed, 98 insertions(+), 77 deletions(-)

diff --git a/theories/admissible.v b/theories/admissible.v
index e7f0769..3e48d49 100644
--- a/theories/admissible.v
+++ b/theories/admissible.v
@@ -120,6 +120,13 @@ Proof.
   hauto l:on use:regularity, E_FunExt.
 Qed.
 
+Lemma E_PairExt Γ (a b : PTm) A B :
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ b ∈ PBind PSig A B ->
+  Γ ⊢ PProj PL a ≡ PProj PL b ∈ A ->
+  Γ ⊢ PProj PR a ≡ PProj PR b ∈ subst_PTm (scons (PProj PL a) VarPTm) B ->
+  Γ ⊢ a ≡ b ∈ PBind PSig A B. hauto l:on use:regularity, E_PairExt. Qed.
+
 Lemma renaming_comp Γ Δ Ξ ξ0 ξ1 :
   renaming_ok Γ Δ ξ0 -> renaming_ok Δ Ξ ξ1 ->
   renaming_ok Γ Ξ (funcomp ξ0 ξ1).
@@ -166,3 +173,92 @@ Proof.
   sauto lq:on use:renaming, renaming_shift, E_Refl.
   constructor. constructor=>//. constructor.
 Qed.
+
+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 : Γ u U / hu => //=.
+  - move => Γ a A B ha _ a0 [*]. subst.
+    exists A, B. split => //=.
+    eapply regularity in ha.
+    move : ha => [i].
+    move /Bind_Inv => [j][h _].
+    by move /E_Refl /Su_Eq in h.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+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 : Γ u U / hu => //=.
+  - move => Γ a A B ha _ a0 [*]. subst.
+    exists A, B. split => //=.
+    have ha' := ha.
+    eapply regularity in ha.
+    move : ha => [i ha].
+    move /T_Proj1 in ha'.
+    apply : bind_inst; eauto.
+    apply : E_Refl ha'.
+  - hauto lq:on rew:off ctrs:LEq.
+Qed.
+
+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 : Γ 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 E_ProjPair1 : forall (a b : PTm) Γ (A : PTm),
+    Γ ⊢ PProj PL (PPair a b) ∈ A -> Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A.
+Proof.
+  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.
+  move /Su_Sig_Proj1 in hS.
+  have {hA0} {}hS : Γ ⊢ A1 ≲ A by eauto using Su_Transitive.
+  apply : E_Conv; eauto.
+  apply : E_ProjPair1; eauto.
+Qed.
+
+Lemma E_ProjPair2 : forall (a b : PTm) Γ (A : PTm),
+    Γ ⊢ PProj PR (PPair a b) ∈ A -> Γ ⊢ PProj PR (PPair a b) ≡ b ∈ A.
+Proof.
+  move => a b Γ A. move /Proj2_Inv.
+  move => [A0][B0][ha]h.
+  have hr := ha.
+  move /Pair_Inv : ha => [A1][B1][ha][hb]hU.
+  have [i hSig] : exists i, Γ ⊢ PBind PSig A1 B1 ∈ PUniv i by sfirstorder use:regularity.
+  have /E_Symmetric : Γ ⊢ (PProj PL (PPair a b)) ≡ a ∈ A1 by eauto using  E_ProjPair1 with wt.
+  move /Su_Sig_Proj2 : hU. repeat move/[apply]. move => hB.
+  apply : E_Conv; eauto.
+  apply : E_Conv; eauto.
+  apply : E_ProjPair2; eauto.
+Qed.
+
+
+Lemma E_PairEta Γ a A B :
+  Γ ⊢ a ∈ PBind PSig A B ->
+  Γ ⊢ PPair (PProj PL a) (PProj PR a) ≡ a ∈ PBind PSig A B.
+Proof.
+  move => h.
+  have [i hSig] : exists i, Γ ⊢ PBind PSig A B ∈ PUniv i by hauto use:regularity.
+  apply E_PairExt => //.
+  eapply T_Pair; eauto with wt.
+  apply : E_Transitive. apply E_ProjPair1. by eauto with wt.
+  by eauto with wt.
+  apply E_ProjPair2.
+  apply : T_Proj2; eauto with wt.
+Qed.
diff --git a/theories/algorithmic.v b/theories/algorithmic.v
index cca7cfa..4e4e6fd 100644
--- a/theories/algorithmic.v
+++ b/theories/algorithmic.v
@@ -608,13 +608,12 @@ 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.
-    hauto l:on use:regularity.
     apply T_Conv with (A := A);eauto.
     eauto using Su_Eq.
     move => ?.
     suff : Γ ⊢ PAbs a ≡ PAbs (PApp (ren_PTm shift u) (VarPTm var_zero)) ∈ PBind PPi A2 B2
       by eauto using E_Transitive.
-    apply : E_Abs; eauto. hauto l:on use:regularity.
+    apply : E_Abs; eauto.
     apply iha.
     move /Su_Pi_Proj2_Var in hPi'.
     apply : T_Conv; eauto.
@@ -666,7 +665,7 @@ Proof.
     have hA02 : Γ ⊢ A0 ≲ A2 by sfirstorder use:Su_Sig_Proj1.
     have hu'  : Γ ⊢ u ∈ PBind PSig A2 B2 by eauto using T_Conv_E.
     move => [:ih0'].
-    apply : E_Transitive; last (apply E_Symmetric; apply : E_PairEta).
+    apply : E_Transitive; last (apply : E_PairEta).
     apply : E_Pair; eauto. hauto l:on use:regularity.
     abstract : ih0'.
     apply ih0. apply : T_Conv; eauto.
@@ -679,7 +678,6 @@ Proof.
     move /E_Symmetric in ih0'.
     move /regularity_sub0 in hA'.
     hauto l:on use:bind_inst.
-    hauto l:on use:regularity.
     eassumption.
   (* Same as before *)
   - move {hAppL}.
diff --git a/theories/preservation.v b/theories/preservation.v
index 2722ee7..301553e 100644
--- a/theories/preservation.v
+++ b/theories/preservation.v
@@ -4,51 +4,6 @@ Require Import ssreflect.
 Require Import Psatz.
 Require Import Coq.Logic.FunctionalExtensionality.
 
-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 : Γ u U / hu => //=.
-  - move => Γ a A B ha _ a0 [*]. subst.
-    exists A, B. split => //=.
-    eapply regularity in ha.
-    move : ha => [i].
-    move /Bind_Inv => [j][h _].
-    by move /E_Refl /Su_Eq in h.
-  - hauto lq:on rew:off ctrs:LEq.
-Qed.
-
-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 : Γ u U / hu => //=.
-  - move => Γ a A B ha _ a0 [*]. subst.
-    exists A, B. split => //=.
-    have ha' := ha.
-    eapply regularity in ha.
-    move : ha => [i ha].
-    move /T_Proj1 in ha'.
-    apply : bind_inst; eauto.
-    apply : E_Refl ha'.
-  - hauto lq:on rew:off ctrs:LEq.
-Qed.
-
-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 : Γ 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 Γ P (a : PTm) b c U :
   Γ ⊢ PInd P a b c ∈ U ->
@@ -91,34 +46,6 @@ Proof.
   asimpl. rewrite !subst_scons_id. by apply E_Symmetric.
 Qed.
 
-Lemma E_ProjPair1 : forall (a b : PTm) Γ (A : PTm),
-    Γ ⊢ PProj PL (PPair a b) ∈ A -> Γ ⊢ PProj PL (PPair a b) ≡ a ∈ A.
-Proof.
-  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.
-  move /Su_Sig_Proj1 in hS.
-  have {hA0} {}hS : Γ ⊢ A1 ≲ A by eauto using Su_Transitive.
-  apply : E_Conv; eauto.
-  apply : E_ProjPair1; eauto.
-Qed.
-
-Lemma E_ProjPair2 : forall (a b : PTm) Γ (A : PTm),
-    Γ ⊢ PProj PR (PPair a b) ∈ A -> Γ ⊢ PProj PR (PPair a b) ≡ b ∈ A.
-Proof.
-  move => a b Γ A. move /Proj2_Inv.
-  move => [A0][B0][ha]h.
-  have hr := ha.
-  move /Pair_Inv : ha => [A1][B1][ha][hb]hU.
-  have [i hSig] : exists i, Γ ⊢ PBind PSig A1 B1 ∈ PUniv i by sfirstorder use:regularity.
-  have /E_Symmetric : Γ ⊢ (PProj PL (PPair a b)) ≡ a ∈ A1 by eauto using  E_ProjPair1 with wt.
-  move /Su_Sig_Proj2 : hU. repeat move/[apply]. move => hB.
-  apply : E_Conv; eauto.
-  apply : E_Conv; eauto.
-  apply : E_ProjPair2; eauto.
-Qed.
-
 Lemma E_Pair Γ a0 b0 a1 b1 A B i :
   Γ ⊢ PBind PSig A B ∈ PUniv i ->
   Γ ⊢ a0 ≡ a1 ∈ A ->