From 0e5b82b162104717571074583491b1c54db2f133 Mon Sep 17 00:00:00 2001 From: Yiyun Liu Date: Thu, 6 Feb 2025 21:40:26 -0500 Subject: [PATCH 01/39] 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 Date: Thu, 6 Feb 2025 23:41:38 -0500 Subject: [PATCH 02/39] 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 Date: Fri, 7 Feb 2025 00:39:34 -0500 Subject: [PATCH 03/39] 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 Date: Fri, 7 Feb 2025 00:43:12 -0500 Subject: [PATCH 04/39] 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 Date: Fri, 7 Feb 2025 01:19:19 -0500 Subject: [PATCH 05/39] 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 Date: Fri, 7 Feb 2025 16:45:58 -0500 Subject: [PATCH 06/39] 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 Date: Sat, 8 Feb 2025 20:37:46 -0500 Subject: [PATCH 07/39] 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 Date: Sat, 8 Feb 2025 21:06:51 -0500 Subject: [PATCH 08/39] 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 Date: Sat, 8 Feb 2025 22:15:04 -0500 Subject: [PATCH 09/39] 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 Date: Sat, 8 Feb 2025 22:38:28 -0500 Subject: [PATCH 10/39] 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 Date: Sat, 8 Feb 2025 22:45:07 -0500 Subject: [PATCH 11/39] 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 Date: Sat, 8 Feb 2025 22:52:50 -0500 Subject: [PATCH 12/39] 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 Date: Sat, 8 Feb 2025 22:56:45 -0500 Subject: [PATCH 13/39] 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 Date: Sun, 9 Feb 2025 16:12:57 -0500 Subject: [PATCH 14/39] 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 Date: Sun, 9 Feb 2025 16:25:34 -0500 Subject: [PATCH 15/39] 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 Date: Sun, 9 Feb 2025 16:33:43 -0500 Subject: [PATCH 16/39] 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 Date: Sun, 9 Feb 2025 16:46:17 -0500 Subject: [PATCH 17/39] 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 Date: Sun, 9 Feb 2025 17:05:36 -0500 Subject: [PATCH 18/39] 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 Date: Sun, 9 Feb 2025 17:05:43 -0500 Subject: [PATCH 19/39] 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 Date: Sun, 9 Feb 2025 20:41:27 -0500 Subject: [PATCH 20/39] 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 Date: Sun, 9 Feb 2025 23:32:07 -0500 Subject: [PATCH 21/39] 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 Date: Mon, 10 Feb 2025 00:17:01 -0500 Subject: [PATCH 22/39] 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 Date: Mon, 10 Feb 2025 01:15:44 -0500 Subject: [PATCH 23/39] 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 Date: Mon, 10 Feb 2025 13:51:35 -0500 Subject: [PATCH 24/39] 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 Date: Mon, 10 Feb 2025 14:16:43 -0500 Subject: [PATCH 25/39] 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 Date: Mon, 10 Feb 2025 17:01:40 -0500 Subject: [PATCH 26/39] 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 Date: Mon, 10 Feb 2025 18:40:42 -0500 Subject: [PATCH 27/39] 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 Date: Mon, 10 Feb 2025 21:50:23 -0500 Subject: [PATCH 28/39] 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 Date: Mon, 10 Feb 2025 21:51:27 -0500 Subject: [PATCH 29/39] 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 Date: Tue, 11 Feb 2025 19:15:06 -0500 Subject: [PATCH 30/39] 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 Date: Wed, 12 Feb 2025 15:54:42 -0500 Subject: [PATCH 31/39] 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 Date: Wed, 12 Feb 2025 15:56:35 -0500 Subject: [PATCH 32/39] 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 Date: Wed, 12 Feb 2025 16:14:51 -0500 Subject: [PATCH 33/39] 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 Date: Wed, 12 Feb 2025 16:40:51 -0500 Subject: [PATCH 34/39] 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 Date: Wed, 12 Feb 2025 19:27:42 -0500 Subject: [PATCH 35/39] 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 Date: Wed, 12 Feb 2025 19:53:20 -0500 Subject: [PATCH 36/39] 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 Date: Wed, 12 Feb 2025 20:18:12 -0500 Subject: [PATCH 37/39] 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 Date: Wed, 12 Feb 2025 21:13:47 -0500 Subject: [PATCH 38/39] 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 Date: Wed, 12 Feb 2025 22:00:44 -0500 Subject: [PATCH 39/39] 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.