Fix the fundamental theorem yet again

This commit is contained in:
Yiyun Liu 2025-02-09 16:46:17 -05:00
parent 5b925e3fa1
commit 20bf38a3ca

View file

@ -702,7 +702,7 @@ Proof.
Qed. 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 -> Γ A PUniv i ->
funcomp (ren_PTm shift) (scons A Γ) B PUniv j -> funcomp (ren_PTm shift) (scons A Γ) B PUniv j ->
Γ PBind p A B PUniv (max i j). Γ PBind p A B PUniv (max i j).
@ -717,6 +717,17 @@ Proof.
- move => *. asimpl. hauto l:on use:InterpUniv_cumulative, ρ_ok_cons. - move => *. asimpl. hauto l:on use:InterpUniv_cumulative, ρ_ok_cons.
Qed. 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 : Lemma ST_Abs n Γ (a : PTm (S n)) A B i :
Γ PBind PPi A B (PUniv i) -> Γ PBind PPi A B (PUniv i) ->
funcomp (ren_PTm shift) (scons A Γ) a B -> 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 Γ)). Γ_eq (funcomp (ren_PTm shift) (scons A Γ)) (funcomp (ren_PTm shift) (scons B Γ)).
Proof. eauto using Γ_eq_refl ,Γ_eq_cons. Qed. 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 -> Γ A0 A1 PUniv i ->
funcomp (ren_PTm shift) (scons A0 Γ) B0 B1 PUniv j -> funcomp (ren_PTm shift) (scons A0 Γ) B0 B1 PUniv j ->
@ -999,8 +1010,8 @@ Proof.
move => hA hB. move => hA hB.
apply SemEq_SemWt in hA, hB. apply SemEq_SemWt in hA, hB.
apply SemWt_SemEq; last by hauto l:on use:DJoin.BindCong. apply SemWt_SemEq; last by hauto l:on use:DJoin.BindCong.
hauto l:on use:ST_Bind. hauto l:on use:ST_Bind'.
apply ST_Bind; first by tauto. apply ST_Bind'; first by tauto.
have hΓ' : funcomp (ren_PTm shift) (scons A1 Γ) by hauto l:on use:SemWff_cons. have hΓ' : funcomp (ren_PTm shift) (scons A1 Γ) by hauto l:on use:SemWff_cons.
move => ρ hρ. move => ρ hρ.
suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on. suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
@ -1008,6 +1019,15 @@ Proof.
hauto lq:on use:Γ_eq_cons'. hauto lq:on use:Γ_eq_cons'.
Qed. 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 : Lemma SE_Abs n Γ (a b : PTm (S n)) A B i :
Γ PBind PPi A B (PUniv i) -> Γ PBind PPi A B (PUniv i) ->
funcomp (ren_PTm shift) (scons A Γ) a b B -> funcomp (ren_PTm shift) (scons A Γ) a b B ->
@ -1234,8 +1254,8 @@ Proof.
move : hA => [hA0][i][hA1]hA2. move : hA => [hA0][i][hA1]hA2.
move : hB => [hB0][j][hB1]hB2. move : hB => [hB0][j][hB1]hB2.
apply : SemWt_SemLEq; last by hauto l:on use:Sub.PiCong. apply : SemWt_SemLEq; last by hauto l:on use:Sub.PiCong.
hauto l:on use:ST_Bind. hauto l:on use:ST_Bind'.
apply ST_Bind; eauto. 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 A1 Γ) by hauto l:on use:SemWff_cons.
move => ρ hρ. move => ρ hρ.
suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on. suff : ρ_ok (funcomp (ren_PTm shift) (scons A0 Γ)) ρ by hauto l:on.
@ -1256,8 +1276,8 @@ Proof.
move : hA => [hA0][i][hA1]hA2. move : hA => [hA0][i][hA1]hA2.
move : hB => [hB0][j][hB1]hB2. move : hB => [hB0][j][hB1]hB2.
apply : SemWt_SemLEq; last by hauto l:on use:Sub.SigCong. apply : SemWt_SemLEq; last by hauto l:on use:Sub.SigCong.
2 : { hauto l:on use:ST_Bind. } 2 : { hauto l:on use:ST_Bind'. }
apply ST_Bind; eauto. 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 A1 Γ) by hauto l:on use:SemWff_cons.
have hΓ'' : funcomp (ren_PTm shift) (scons A0 Γ) by hauto l:on use:SemWff_cons. have hΓ'' : funcomp (ren_PTm shift) (scons A0 Γ) by hauto l:on use:SemWff_cons.
move => ρ hρ. move => ρ hρ.