Compare commits

..

39 commits

Author SHA1 Message Date
Yiyun Liu
1f1b8a83db Pull out some inversion lemmas to prove later 2025-02-12 22:00:44 -05:00
Yiyun Liu
1d3920fce1 Prove the case for pair and neutral 2025-02-12 21:13:47 -05:00
Yiyun Liu
ba77752329 Add more cases to the soundness proof 2025-02-12 20:18:12 -05:00
Yiyun Liu
48adb34946 Add simplified projection lemma 2025-02-12 19:53:20 -05:00
Yiyun Liu
d053f93100 Finish the app neutral case 2025-02-12 19:27:42 -05:00
fa80294c5d Minor 2025-02-12 16:40:51 -05:00
761e96f414 Use the abstract tactic to finish off the symmetric casesa 2025-02-12 16:14:51 -05:00
5ac2bf1c40 Minor 2025-02-12 15:56:35 -05:00
823f61d89f Finish most cases of the soundness proof 2025-02-12 15:54:42 -05:00
15f8a9c687 Try a few cases of soundness 2025-02-11 19:15:06 -05:00
c1a8e9d2e1 Add the top-level subject reduction theorem 2025-02-10 21:51:27 -05:00
c5de86339f Finish subject reduction 2025-02-10 21:50:23 -05:00
bccf6eb860 Add Coquand's algorithm 2025-02-10 18:40:42 -05:00
8105b5c410 Add admissible simple rules 2025-02-10 17:01:40 -05:00
8f8f428562 Finish preservation 2025-02-10 14:16:43 -05:00
5918fdf47e Prove all but 5 cases of regularity 2025-02-10 13:51:35 -05:00
Yiyun Liu
26e3c1c42a Add some cases for regularity 2025-02-10 01:15:44 -05:00
Yiyun Liu
c8e84ef93c Finish morphing 2025-02-10 00:17:01 -05:00
Yiyun Liu
2c47d78ad6 Add stub for morphing 2025-02-09 23:32:07 -05:00
Yiyun Liu
ea1fba8ae9 Finish syntactic renaming 2025-02-09 20:41:27 -05:00
881b2e3825 Add missing premise 2025-02-09 17:05:43 -05:00
df5c6bf0b1 Minor 2025-02-09 17:05:36 -05:00
20bf38a3ca Fix the fundamental theorem yet again 2025-02-09 16:46:17 -05:00
5b925e3fa1 Add beta and eta rules to syntactic typing 2025-02-09 16:33:43 -05:00
133968dd23 Add semantic eta laws for pair 2025-02-09 16:25:34 -05:00
ab1bd8eef8 Add semantic rules for function beta and eta 2025-02-09 16:12:57 -05:00
Yiyun Liu
4396786701 Reformulate renaming 2025-02-08 22:56:45 -05:00
Yiyun Liu
932662d5d9 Finish soundness proof 2025-02-08 22:52:50 -05:00
Yiyun Liu
0c83044d72 Update the typing rules with projections 2025-02-08 22:45:07 -05:00
Yiyun Liu
5996c45526 Finish the semantic projection rules 2025-02-08 22:38:28 -05:00
Yiyun Liu
02e6c9e025 Add pi and sig subtyping semantic rules 2025-02-08 22:15:04 -05:00
Yiyun Liu
916e0bcd75 Change the conversion rules to use Sub instead of Eq 2025-02-08 21:06:51 -05:00
Yiyun Liu
f483d63f01 Fix the definition of semleq 2025-02-08 20:37:46 -05:00
0746e9a354 Finish subtyping semantics 2025-02-07 16:45:58 -05:00
4bc08e1897 Add one interpuniv sub case 2025-02-07 01:19:19 -05:00
707483d401 Add injectivity for subtyping 2025-02-07 00:43:12 -05:00
2f4ea2c78b Add more noconfusion lemmas for untyped subtyping 2025-02-07 00:39:34 -05:00
cf2726be8d Add subtyping 2025-02-06 23:41:38 -05:00
0e5b82b162 Move projection axioms to the subtyping relation 2025-02-06 21:40:26 -05:00
9 changed files with 2249 additions and 232 deletions

69
theories/admissible.v Normal file
View file

@ -0,0 +1,69 @@
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 : funcomp (ren_PTm shift) (scons A Γ) by sfirstorder use:wff_mutual.
move /Wff_Cons_Inv : => [][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.
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.

386
theories/algorithmic.v Normal file
View file

@ -0,0 +1,386 @@
Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax
common typing preservation admissible fp_red structural.
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).
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.
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.
(* 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).
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_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 ->
(* --------------------- *)
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
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) and "a b" := (CoqEq_Neu a b).
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_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 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) /\
(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 => [ 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 [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;
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'.
move => [A1][B1][hu1']hu1''.
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 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).
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.
apply E_Refl. eauto using T_Proj1.
apply : Su_Transitive ;eauto.
apply : Su_Sig_Proj2; eauto.
apply : E_Proj1; eauto.
apply : E_Proj2; eauto.
- move => n u0 u1 a0 a1 neu0 neu1 hu ihu ha iha Γ A B wta0 wta1.
move /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 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
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_Transitive; eauto.
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'.
have [i [A2 [B2 h]]] : exists i A2 B2, Γ A PBind PPi A2 B2 PUniv i by hauto l:on use:Sub_Bind_InvR.
have hp0 : Γ PBind PPi A0 B0 PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
have hp1 : Γ PBind PPi A1 B1 PBind PPi A2 B2 by eauto using Su_Transitive, Su_Eq.
have [j ?] : exists j, Γ A0 PUniv j by qauto l:on use:Wff_Cons_Inv, wff_mutual.
have [k ?] : exists j, Γ A1 PUniv j by qauto l:on use:Wff_Cons_Inv, wff_mutual.
have [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.
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.
move /Su_Pi_Proj2_Var in hp1.
apply : T_Conv; eauto.
eapply ctx_eq_subst_one with (A0 := A1); eauto.
- abstract : hAppL.
move => n a u hneu ha iha Γ A wta wtu.
move /Abs_Inv : wta => [A0][B0][wta]hPi.
have [i [A2 [B2 h]]] : exists i A2 B2, Γ A PBind PPi A2 B2 PUniv i by hauto l:on use:Sub_Bind_InvR.
have hPi'' : Γ PBind PPi A2 B2 A by eauto using Su_Eq, Su_Transitive, E_Symmetric.
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.
apply : E_AppEta; eauto.
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:regularity.
apply iha.
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);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.
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 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.
have /Su_Sig_Proj1 h0' := h0.
have /Su_Sig_Proj1 h1' := h1.
move => [:eqa].
apply : E_Pair; eauto. hauto l:on use:regularity.
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 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.
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.
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.
eassumption.
(* Same as before *)
- move {hAppL}.
move => *. apply E_Symmetric. apply : hPairL;
sfirstorder use:coqeq_symmetric_mutual, E_Symmetric.
- sfirstorder use:E_Refl.
- move => {hAppL hPairL} n p A0 A1 B0 B1 hA ihA hB ihB Γ A hA0 hA1.
move /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 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.
apply T_Conv with (A := PUniv j); eauto.
apply ihB.
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.
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.
Qed.

92
theories/common.v Normal file
View file

@ -0,0 +1,92 @@
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].
apply iha in h. by subst.
destruct i, j=>//=.
hauto l:on.
move => n p A ihA B ihB m ξ []//=.
move => b A0 B0 [?]. subst.
move => ?. have ? : A0 = A by firstorder. subst.
move => ?. have : B = B0. apply : ihB; eauto.
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.

View file

@ -8,7 +8,7 @@ Require Import Arith.Wf_nat (well_founded_lt_compat).
Require Import Psatz. Require Import Psatz.
From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn). From stdpp Require Import relations (rtc (..), rtc_once, rtc_r, sn).
From Hammer Require Import Tactics. 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 Btauto.
Require Import Cdcl.Itauto. Require Import Cdcl.Itauto.
@ -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. 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 : Lemma PProj_imp n p a :
@ishf n a -> @ishf n a ->
@ -1408,35 +1373,6 @@ Module ERed.
(* destruct a. *) (* destruct a. *)
(* exact (ξ f). *) (* 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].
apply iha in h. by subst.
destruct i, j=>//=.
hauto l:on.
move => n p A ihA B ihB m ξ []//=.
move => b A0 B0 [?]. subst.
move => ?. have ? : A0 = A by firstorder. subst.
move => ?. have : B = B0. apply : ihB; eauto.
sauto.
congruence.
Qed.
Lemma AppEta' n a u : Lemma AppEta' n a u :
u = (@PApp (S n) (ren_PTm shift a) (VarPTm var_zero)) -> u = (@PApp (S n) (ren_PTm shift a) (VarPTm var_zero)) ->
R (PAbs u) a. R (PAbs u) a.
@ -1774,6 +1710,17 @@ Module REReds.
eauto using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl. eauto using AppCong, AbsCong, BindCong, ProjCong, PairCong, cong_up, rtc_refl.
Qed. 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. End REReds.
Module LoRed. Module LoRed.
@ -2286,3 +2233,242 @@ Module DJoin.
Qed. Qed.
End DJoin. 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)
| PiCong A0 A1 B0 B1 :
R A1 A0 ->
R B0 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).
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; 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.
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.
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.
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 PiCong n (A0 A1 : PTm n) B0 B1 :
R A1 A0 ->
R B0 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 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 :
i <= j ->
@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 b -> isuniv a -> 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 bind_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
R (PBind p0 A0 B0) (PBind p1 A1 B1) ->
p0 = p1 /\ (if p0 is PPi then R A1 A0 else R A0 A1) /\ R B0 B1.
Proof.
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; 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.
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.

View file

@ -1,5 +1,5 @@
Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax. Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
Require Import fp_red. Require Import common fp_red.
From Hammer Require Import Tactics. From Hammer Require Import Tactics.
From Equations Require Import Equations. From Equations Require Import Equations.
Require Import ssreflect ssrbool. Require Import ssreflect ssrbool.
@ -24,6 +24,7 @@ Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop)
| InterpExt_Ne A : | InterpExt_Ne A :
SNe A -> SNe A ->
A i ;; I (fun a => exists v, rtc TRedSN a v /\ SNe v) A i ;; I (fun a => exists v, rtc TRedSN a v /\ SNe v)
| InterpExt_Bind p A B PA PF : | InterpExt_Bind p A B PA PF :
A i ;; I PA -> A i ;; I PA ->
(forall a, PA a -> exists PB, PF a PB) -> (forall a, PA a -> exists PB, PF a PB) ->
@ -40,7 +41,6 @@ Inductive InterpExt {n} i (I : nat -> PTm n -> Prop) : PTm n -> (PTm n -> Prop)
A i ;; I PA A i ;; I PA
where "⟦ A ⟧ i ;; I ↘ S" := (InterpExt i I A S). where "⟦ A ⟧ i ;; I ↘ S" := (InterpExt i I A S).
Lemma InterpExt_Univ' n i I j (PF : PTm n -> Prop) : Lemma InterpExt_Univ' n i I j (PF : PTm n -> Prop) :
PF = I j -> PF = I j ->
j < i -> j < i ->
@ -261,7 +261,8 @@ Qed.
Lemma redsns_preservation : forall n a b, @SN n a -> rtc TRedSN a b -> SN b. 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. 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 : Lemma InterpUniv_SNe_inv n i (A : PTm n) PA :
SNe A -> SNe A ->
@ -294,26 +295,150 @@ Proof.
subst. hauto lq:on inv:TRedSN. subst. hauto lq:on inv:TRedSN.
Qed. Qed.
Lemma bindspace_iff n p (PA : PTm n -> Prop) PF PF0 b : Lemma bindspace_impl n p (PA PA0 : PTm n -> Prop) PF PF0 b :
(forall (a : PTm n) (PB PB0 : PTm n -> Prop), PA a -> PF a PB -> PF0 a PB0 -> PB = PB0) -> (forall 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 PB, PF a PB) ->
(forall a, PA a -> exists PB0, PF0 a PB0) -> (forall a, PA0 a -> exists PB0, PF0 a PB0) ->
(BindSpace p PA PF b <-> BindSpace p PA PF0 b). (BindSpace p PA PF b -> BindSpace p PA0 PF0 b).
Proof. Proof.
rewrite /BindSpace => h hPF hPF0. rewrite /BindSpace => hSA h hPF hPF0.
case : p => /=. case : p hSA => /= hSA.
- rewrite /ProdSpace. - rewrite /ProdSpace.
split.
move => h1 a PB ha hPF'. move => h1 a PB ha hPF'.
specialize hPF with (1 := ha). have {}/hPF : PA a by sfirstorder.
specialize hPF0 with (1 := ha). specialize hPF0 with (1 := ha).
sblast. hauto lq:on.
move => ? a PB ha.
specialize hPF with (1 := ha).
specialize hPF0 with (1 := ha).
sblast.
- rewrite /SumSpace. - 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 :
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. 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
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. hauto l: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'.
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_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.
Proof.
move : InterpUniv_Sub'. repeat move/[apply].
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. Qed.
Lemma InterpUniv_Join n i (A B : PTm n) PA PB : Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
@ -322,62 +447,19 @@ Lemma InterpUniv_Join n i (A B : PTm n) PA PB :
DJoin.R A B -> DJoin.R A B ->
PA = PB. PA = PB.
Proof. Proof.
move => hA. move => + + /[dup] /Sub.FromJoin + /DJoin.symmetric /Sub.FromJoin.
move : i A PA hA B PB. move : InterpUniv_Sub'; repeat move/[apply]. move => h.
apply : InterpUniv_ind. move => h1 h2.
- move => i A hA B PB hPB hAB. extensionality x.
have [*] : SN B /\ SN A by hauto l:on use:adequacy. apply propositional_extensionality.
move /InterpUniv_case : hPB. firstorder.
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.
Qed. Qed.
Lemma InterpUniv_Functional n i (A : PTm n) PA PB : Lemma InterpUniv_Functional n i (A : PTm n) PA PB :
A i PA -> A i PA ->
A i PB -> A i PB ->
PA = 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 : Lemma InterpUniv_Join' n i j (A B : PTm n) PA PB :
A i PA -> A i PA ->
@ -431,17 +513,36 @@ Qed.
Definition ρ_ok {n} (Γ : fin n -> PTm n) (ρ : fin n -> PTm 0) := forall i k PA, Definition ρ_ok {n} (Γ : fin n -> PTm n) (ρ : fin n -> PTm 0) := forall i k PA,
subst_PTm ρ (Γ i) k PA -> PA (ρ i). 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). 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). Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70).
Definition SemEq {n} Γ (a b A : PTm n) := DJoin.R a b /\ forall ρ, ρ_ok Γ ρ -> exists k PA, subst_PTm ρ A k PA /\ PA (subst_PTm ρ a) /\ PA (subst_PTm ρ b). Definition SemEq {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). Notation "Γ ⊨ a ≡ b ∈ A" := (SemEq Γ a b A) (at level 70).
Definition SemLEq {n} Γ (A B : PTm n) := Sub.R A B /\ exists i, forall ρ, ρ_ok Γ ρ -> exists S0 S1, subst_PTm ρ A i S0 /\ subst_PTm ρ B i S1.
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. 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. 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. Lemma SemWt_SemEq n Γ (a b A : PTm n) : Γ a A -> Γ b A -> (DJoin.R a b) -> Γ a b A.
Proof. Proof.
move => ha hb heq. split => //= ρ hρ. move => ha hb heq. split => //= ρ hρ.
@ -453,6 +554,23 @@ Proof.
hauto lq:on. hauto lq:on.
Qed. Qed.
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 *) (* Semantic context wellformedness *)
Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ Γ i PUniv j. Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ Γ i PUniv j.
Notation "⊨ Γ" := (SemWff Γ) (at level 70). Notation "⊨ Γ" := (SemWff Γ) (at level 70).
@ -488,9 +606,6 @@ Lemma ρ_ok_cons' n i (Γ : fin n -> PTm n) ρ a PA A Δ :
ρ_ok Δ (scons a ρ). ρ_ok Δ (scons a ρ).
Proof. move => ->. apply ρ_ok_cons. Qed. 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) ρ : Lemma ρ_ok_renaming n m (Γ : fin n -> PTm n) ρ :
forall (Δ : fin m -> PTm m) ξ, forall (Δ : fin m -> PTm m) ξ,
renaming_ok Γ Δ ξ -> renaming_ok Γ Δ ξ ->
@ -544,20 +659,10 @@ Lemma SemEq_SN_Join n Γ (a b A : PTm n) :
SN a /\ SN b /\ SN A /\ DJoin.R a b. SN a /\ SN b /\ SN A /\ DJoin.R a b.
Proof. hauto l:on use:SemEq_SemWt, SemWt_SN. Qed. Proof. hauto l:on use:SemEq_SemWt, SemWt_SN. Qed.
Lemma SemWt_Univ n Γ (A : PTm n) i : Lemma SemLEq_SN_Sub n Γ (a b : PTm n) :
Γ A PUniv i <-> Γ a b ->
forall ρ, ρ_ok Γ ρ -> exists S, subst_PTm ρ A i S. SN a /\ SN b /\ Sub.R a b.
Proof. Proof. hauto l:on use:SemLEq_SemWt, SemWt_SN. Qed.
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 *) (* Structural laws for Semantic context wellformedness *)
Lemma SemWff_nil : SemWff null. Lemma SemWff_nil : SemWff null.
@ -597,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).
@ -612,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 ->
@ -650,6 +766,13 @@ Proof.
asimpl. hauto lq:on. asimpl. hauto lq:on.
Qed. 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 : Lemma ST_Pair n Γ (a b : PTm n) A B i :
Γ PBind PSig A B (PUniv i) -> Γ PBind PSig A B (PUniv i) ->
Γ a A -> Γ a A ->
@ -751,23 +874,34 @@ Qed.
Lemma ST_Conv' n Γ (a : PTm n) A B i : Lemma ST_Conv' n Γ (a : PTm n) A B i :
Γ a A -> Γ a A ->
Γ B PUniv i -> Γ B PUniv i ->
DJoin.R A B -> Sub.R A B ->
Γ a B. Γ a B.
Proof. Proof.
move => ha /SemWt_Univ h h0. move => ha /SemWt_Univ h h0.
move => ρ hρ. 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 /ha : (hρ){ha} => [m [PA [h1 h2]]].
move /h : (hρ){h} => [S hS]. move /h : (hρ){h} => [S hS].
have ? : PA = S by eauto using InterpUniv_Join'. subst. have h3 : forall x, PA x -> S x.
eauto. move : InterpUniv_Sub h0 h1 hS; by repeat move/[apply].
hauto lq:on.
Qed. 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 A ->
Γ A B PUniv i -> Γ B PUniv i ->
DJoin.R A B ->
Γ 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 : Lemma SE_Refl n Γ (a : PTm n) A :
Γ a A -> Γ a A ->
@ -791,6 +925,8 @@ Proof.
hauto l:on use:DJoin.transitive. hauto l:on use:DJoin.transitive.
Qed. 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 Δ ρ. Lemma Γ_eq_ρ_ok n Γ Δ (ρ : fin n -> PTm 0) : Γ_eq Γ Δ -> Γ -> ρ_ok Γ ρ -> ρ_ok Δ ρ.
Proof. Proof.
move => hΓΔ h. move => hΓΔ h.
@ -806,6 +942,44 @@ Proof.
hauto l:on use: DJoin.substing. hauto l:on use: DJoin.substing.
Qed. 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.
move => i k PA hPA.
move : . rewrite /SemWff. move /(_ i) => [j].
move => .
rewrite SemWt_Univ in .
have {}/ := 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) : Lemma Γ_eq_refl n (Γ : fin n -> PTm n) :
Γ_eq Γ Γ. Γ_eq Γ Γ.
Proof. sfirstorder use:DJoin.refl. Qed. Proof. sfirstorder use:DJoin.refl. Qed.
@ -822,13 +996,12 @@ Proof.
rewrite /funcomp. rewrite /funcomp.
asimpl. substify. apply DJoin.substing. by asimpl. asimpl. substify. apply DJoin.substing. by asimpl.
Qed. Qed.
Lemma Γ_eq_cons' n (Γ : fin n -> PTm n) A B : Lemma Γ_eq_cons' n (Γ : fin n -> PTm n) A B :
DJoin.R A B -> DJoin.R 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 ->
@ -837,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.
@ -846,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 ->
@ -855,22 +1037,66 @@ Proof.
apply SemWt_SemEq; eauto using DJoin.AbsCong, ST_Abs. apply SemWt_SemEq; eauto using DJoin.AbsCong, ST_Abs.
Qed. 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)) b PBind PPi A B.
Proof.
move => h0 h1. apply SemWt_SemEq; eauto.
apply : ST_Abs; eauto.
have hA : Γ A PUniv i by eauto using SBind_inv1.
eapply ST_App' with (A := ren_PTm shift A)(B:= ren_PTm (upRen_PTm_PTm shift) B). by asimpl.
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.
hauto q:on ctrs:rtc,RERed.R.
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 : Lemma SE_Conv' n Γ (a b : PTm n) A B i :
Γ a b A -> Γ a b A ->
Γ B PUniv i -> Γ B PUniv i ->
DJoin.R A B -> Sub.R A B ->
Γ a b B. Γ a b B.
Proof. Proof.
move /SemEq_SemWt => [ha][hb]he hB hAB. move /SemEq_SemWt => [ha][hb]he hB hAB.
apply SemWt_SemEq; eauto using ST_Conv'. apply SemWt_SemEq; eauto using ST_Conv'.
Qed. 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 A ->
Γ A B PUniv i -> Γ A B ->
Γ a b B. Γ a b B.
Proof. Proof.
move => h /SemEq_SemWt [h0][h1]he. move => h /SemLEq_SemWt [h0][h1][ha]hb.
eauto using SE_Conv'. eauto using SE_Conv'.
Qed. Qed.
@ -901,7 +1127,7 @@ Lemma SE_Pair n Γ (a0 a1 b0 b1 : PTm n) A B i :
Γ PPair a0 b0 PPair a1 b1 PBind PSig A B. Γ PPair a0 b0 PPair a1 b1 PBind PSig A B.
Proof. Proof.
move => h /SemEq_SemWt [ha0][ha1]hae /SemEq_SemWt [hb0][hb1]hbe. 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. Qed.
Lemma SE_Proj1 n Γ (a b : PTm n) A B : Lemma SE_Proj1 n Γ (a b : PTm n) A B :
@ -921,12 +1147,47 @@ Proof.
move => /SemEq_SemWt [ha][hb]he. move => /SemEq_SemWt [ha][hb]he.
apply SemWt_SemEq; eauto using DJoin.ProjCong, ST_Proj2. 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. 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. apply : SBind_inst. eauto using ST_Proj1.
eauto. eauto.
hauto lq:on use: DJoin.cong, DJoin.ProjCong. hauto lq:on use: DJoin.cong, DJoin.ProjCong.
Qed. 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 : Lemma SE_App n Γ i (b0 b1 a0 a1 : PTm n) A B :
Γ PBind PPi A B (PUniv i) -> Γ PBind PPi A B (PUniv i) ->
Γ b0 b1 PBind PPi A B -> Γ b0 b1 PBind PPi A B ->
@ -936,49 +1197,29 @@ Proof.
move => hPi. move => hPi.
move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha. move => /SemEq_SemWt [hb0][hb1]hb /SemEq_SemWt [ha0][ha1]ha.
apply SemWt_SemEq; eauto using DJoin.AppCong, ST_App. 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. Qed.
Lemma SBind_inv1 n Γ i p (A : PTm n) B : Lemma SSu_Eq n Γ (A B : PTm n) i :
Γ PBind p A B PUniv i -> Γ A B PUniv i ->
Γ A PUniv i. Γ A B.
move /SemWt_Univ => h. apply SemWt_Univ. Proof. move /SemEq_SemWt => h.
hauto lq:on rew:off use:InterpUniv_Bind_inv. qauto l:on use:SemWt_SemLEq, Sub.FromJoin.
Qed. Qed.
Lemma SE_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i : Lemma SSu_Transitive n Γ (A B C : PTm n) :
Γ PBind p A0 B0 PBind p A1 B1 PUniv i -> Γ A B ->
Γ A0 A1 PUniv i. Γ B C ->
Γ A C.
Proof. Proof.
move /SemEq_SemWt => [h0][h1]he. move => ha hb.
apply SemWt_SemEq; eauto using SBind_inv1. apply SemLEq_SemWt in ha, hb.
move /DJoin.bind_inj : he. tauto. 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. Qed.
Lemma SE_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i : Lemma ST_Univ' n Γ i j :
Γ 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.
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'; eauto. hauto l:on use:SBind_inv1.
rewrite /DJoin.R.
eexists. split.
apply : relations.rtc_transitive.
apply REReds.substing. apply h2.
apply REReds.cong.
have : forall i0, rtc RERed.R (scons a0 VarPTm i0) (scons c VarPTm i0) by hauto q:on ctrs:rtc inv:option.
apply.
apply : relations.rtc_transitive. apply REReds.substing; eauto.
apply REReds.cong.
hauto lq:on ctrs:rtc inv:option.
Qed.
Lemma ST_Univ n Γ i j :
i < j -> i < j ->
Γ PUniv i : PTm n PUniv j. Γ PUniv i : PTm n PUniv j.
Proof. Proof.
@ -986,6 +1227,109 @@ Proof.
apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ. apply SemWt_Univ. move => ρ hρ. eexists. by apply InterpUniv_Univ.
Qed. 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.
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 => 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 => 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 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 /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 #[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_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 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.

179
theories/preservation.v Normal file
View file

@ -0,0 +1,179 @@
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 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.
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.

View file

@ -5,7 +5,8 @@ From Hammer Require Import Tactics.
Theorem fundamental_theorem : Theorem fundamental_theorem :
(forall n (Γ : fin n -> PTm n), Γ -> Γ) /\ (forall n (Γ : fin n -> PTm n), Γ -> Γ) /\
(forall n Γ (a A : PTm n), Γ a A -> Γ a A) /\ (forall n Γ (a A : PTm n), Γ a A -> Γ a A) /\
(forall n Γ (a b A : PTm n), Γ a b A -> Γ a b A). (forall n Γ (a b A : PTm n), Γ a b A -> Γ a b A) /\
apply wt_mutual; eauto with sem;[idtac]. (forall n Γ (a b : PTm n), Γ a b -> Γ a b).
hauto l:on use:SE_Pair. apply wt_mutual; eauto with sem; [hauto l:on use:SE_Pair].
Unshelve. all : exact 0.
Qed. Qed.

676
theories/structural.v Normal file
View file

@ -0,0 +1,676 @@
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) /\
(forall n Γ (a A : PTm n), Γ a A -> Γ) /\
(forall n Γ (a b A : PTm n), Γ a b A -> Γ) /\
(forall n Γ (A B : PTm n), Γ A B -> Γ).
Proof. apply wt_mutual; 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 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 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 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 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 ->
Γ 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_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 ->
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 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 Δ Γ ξ ->
Δ 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) /\
(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 _ m Δ ξ .
rewrite .
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 Δ ξ .
apply : T_Abs; eauto.
move : ihP() (); 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 Δ ξ .
eapply T_Pair' with (U := ren_PTm ξ (subst_PTm (scons a VarPTm) B));eauto. by asimpl.
- move => n Γ a A B ha iha m Δ ξ . apply : T_Proj2'; eauto. by asimpl.
- hauto lq:on ctrs:Wt,LEq.
- hauto lq:on ctrs:Eq.
- hauto lq:on rew:off use:E_Bind', Wff_Cons, renaming_up.
- move => n Γ a b A B i hPi ihPi ha iha m Δ ξ .
move : ihPi () (). repeat move/[apply].
move => /Bind_Inv [j][h0][h1]h2.
have ? : Δ PBind PPi (ren_PTm ξ A) (ren_PTm (upRen_PTm_PTm ξ) B) PUniv j by qauto l:on ctrs:Wt.
move {hPi}.
apply : E_Abs; eauto. qauto l:on ctrs:Wff use:renaming_up.
- move => *. apply : E_App'; eauto. by asimpl.
- move => n Γ a0 a1 b0 b1 A B i hA ihA ha iha hb ihb m Δ ξ .
apply : E_Pair; eauto.
move : ihb . 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 Δ ξ .
move : ihP () (). repeat move/[apply].
move /Bind_Inv.
move => [j][h0][h1]h2.
have ? : Δ PBind PPi (ren_PTm ξ A) (ren_PTm (upRen_PTm_PTm ξ) B) PUniv j by qauto l:on ctrs:Wt.
apply : E_AppAbs'; eauto. by asimpl. by asimpl.
hauto lq:on ctrs:Wff use:renaming_up.
- move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ξ .
move : {hP} ihP () (). repeat move/[apply].
move /Bind_Inv => [i0][h0][h1]h2.
have ? : Δ PBind PSig (ren_PTm ξ A) (ren_PTm (upRen_PTm_PTm ξ) B) PUniv i0 by qauto l:on ctrs:Wt.
apply : E_ProjPair1; eauto.
move : ihb . repeat move/[apply]. by asimpl.
- move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ξ .
apply : E_ProjPair2'; eauto. by asimpl.
move : ihb ; 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.
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ρ.
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 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 Δ Γ ρ ->
Δ 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 => //=.
- hauto lq:on use:morphing_up, Wff_Cons', T_Bind.
- move => n Γ a A B i hP ihP ha iha m Δ ρ hρ.
move : ihP () (hρ); repeat move/[apply].
move /Bind_Inv => [j][h0][h1]h2. move {hP}.
have ? : Δ PBind PPi (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) PUniv i by hauto lq:on ctrs:Wt.
apply : T_Abs; eauto.
apply : iha.
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ρ .
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ρ.
move : ihPi () (hρ). repeat move/[apply].
move => /Bind_Inv [j][h0][h1]h2.
have ? : Δ PBind PPi (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) PUniv j by qauto l:on ctrs:Wt.
move {hPi}.
apply : E_Abs; eauto. qauto l:on ctrs:Wff use:morphing_up.
- move => *. apply : E_App'; eauto. by asimpl.
- move => n Γ a0 a1 b0 b1 A B i hA ihA ha iha hb ihb m Δ ρ hρ.
apply : E_Pair; eauto.
move : ihb 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ρ.
move : ihP (hρ) (). repeat move/[apply].
move /Bind_Inv.
move => [j][h0][h1]h2.
have ? : Δ PBind PPi (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) PUniv j by qauto l:on ctrs:Wt.
apply : E_AppAbs'; eauto. by asimpl. by asimpl.
hauto lq:on ctrs:Wff use:morphing_up.
- move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ρ hρ.
move : {hP} ihP (hρ) (). repeat move/[apply].
move /Bind_Inv => [i0][h0][h1]h2.
have ? : Δ PBind PSig (subst_PTm ρ A) (subst_PTm (up_PTm_PTm ρ) B) PUniv i0 by qauto l:on ctrs:Wt.
apply : E_ProjPair1; eauto.
move : ihb hρ . repeat move/[apply]. by asimpl.
- move => n Γ a b A B i hP ihP ha iha hb ihb m Δ ρ hρ.
apply : E_ProjPair2'; eauto. by asimpl.
move : ihb 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 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 Γ .
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 [:]. apply : morphing_wt; eauto.
abstract : . sfirstorder use:wff_mutual.
apply morphing_ext; last by asimpl.
by apply morphing_id.
Qed.
(* Could generalize to all equal contexts *)
Lemma ctx_eq_subst_one n (A0 A1 : PTm n) i j Γ a A :
funcomp (ren_PTm shift) (scons A0 Γ) a A ->
Γ 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 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) /\
(forall n Γ (a b A : PTm n), Γ a b A -> Γ a A /\ Γ b A /\ exists i, Γ A PUniv i) /\
(forall n Γ (A B : PTm n), Γ A B -> exists i, Γ A PUniv i /\ Γ B PUniv i).
Proof.
apply wt_mutual => //=; eauto with wt.
- move => n Γ A i ihΓ hA _ j.
destruct j as [j|].
have := ihΓ j.
move => [j0 hj].
exists j0. apply : renaming_wt' => //=; eauto using renaming_shift.
reflexivity. econstructor; eauto.
exists i. rewrite {2}/funcomp. simpl.
apply : renaming_wt'; eauto. reflexivity.
econstructor; eauto.
apply : renaming_shift; eauto.
- move => n Γ b a A B hb [i ihb] ha [j iha].
move /Bind_Inv : ihb => [k][h0][h1]h2.
move : substing_wt ha h1; repeat move/[apply].
move => h. exists k.
move : h. by asimpl.
- hauto lq:on use:Bind_Inv.
- move => n Γ a A B ha [i /Bind_Inv[j][h0][h1]h2].
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 ihΓ hA0
[i0 ihA0] hA [ihA [ihA' [i1 ihA'']]].
move => hB [ihB0 [ihB1 [i2 ihB2]]].
repeat split => //=.
qauto use:T_Bind.
apply T_Bind; eauto.
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]]]
ha [iha0 [iha1 [i1 iha2]]].
repeat split.
qauto use:T_App.
apply : T_Conv; eauto.
qauto use:T_App.
move /E_Symmetric in ha.
by eauto using bind_inst.
hauto lq:on ctrs:Wt,Eq,LEq lq:on use:Bind_Inv, substing_wt.
- hauto lq:on use:bind_inst db:wt.
- hauto lq:on use:Bind_Inv db:wt.
- move => 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: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 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, Bind_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.
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 *. exists (S (max i j)).
have [? ?] : S i <= S (Nat.max i j) /\ S j <= S (Nat.max i j) by lia.
hauto lq:on ctrs:Wt,LEq.
- move => n Γ A0 A1 B0 B1 i 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 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.
- move => n Γ A0 A1 B0 B1 _ [i][ih0 ih1].
move /Bind_Inv : ih0 => [i0][h _].
move /Bind_Inv : ih1 => [i1][h' _].
exists (max i0 i1).
have [? ?] : i0 <= Nat.max i0 i1 /\ i1 <= Nat.max i0 i1 by lia.
eauto using Cumulativity.
- move => n Γ A0 A1 B0 B1 _ [i][ih0 ih1].
move /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 /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.
+ 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 /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.
+ 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.
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 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.
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.
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.

View file

@ -2,16 +2,17 @@ Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax.
Reserved Notation "Γ ⊢ a ∈ A" (at level 70). Reserved Notation "Γ ⊢ a ∈ A" (at level 70).
Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70). Reserved Notation "Γ ⊢ a ≡ b ∈ A" (at level 70).
Reserved Notation "Γ ⊢ A ≲ B" (at level 70).
Reserved Notation "⊢ Γ" (at level 70). Reserved Notation "⊢ Γ" (at level 70).
Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop := Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
| T_Var n Γ (i : fin n) : | T_Var n Γ (i : fin n) :
Γ -> Γ ->
Γ VarPTm i Γ i Γ 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 -> Γ A PUniv i ->
funcomp (ren_PTm shift) (scons A Γ) B PUniv j -> funcomp (ren_PTm shift) (scons A Γ) B PUniv i ->
Γ PBind p A B PUniv (max i j) Γ PBind p A B PUniv i
| T_Abs n Γ (a : PTm (S n)) A B i : | T_Abs n Γ (a : PTm (S n)) A B i :
Γ PBind PPi A B (PUniv i) -> Γ PBind PPi A B (PUniv i) ->
@ -37,16 +38,17 @@ Inductive Wt : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> Prop :=
Γ a PBind PSig A B -> Γ a PBind PSig A B ->
Γ PProj PR a subst_PTm (scons (PProj PL a) VarPTm) B Γ PProj PR a subst_PTm (scons (PProj PL a) VarPTm) B
| T_Univ n Γ i j : | T_Univ n Γ i :
i < j -> Γ ->
Γ PUniv i : PTm n PUniv j Γ 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 A ->
Γ A B PUniv i -> Γ A B ->
Γ a B Γ a B
with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop := with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
(* Structural *)
| E_Refl n Γ (a : PTm n) A : | E_Refl n Γ (a : PTm n) A :
Γ a A -> Γ a A ->
Γ a a A Γ a a A
@ -60,11 +62,13 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
Γ b c A -> Γ b c A ->
Γ a c A Γ a c A
| E_Bind n Γ i j p (A0 A1 : PTm n) B0 B1 : (* Congruence *)
| E_Bind n Γ i p (A0 A1 : PTm n) B0 B1 :
Γ -> Γ ->
Γ A0 PUniv i ->
Γ 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 i ->
Γ PBind p A0 B0 PBind p A1 B1 PUniv (max i j) Γ PBind p A0 B0 PBind p A1 B1 PUniv i
| E_Abs n Γ (a b : PTm (S n)) A B i : | E_Abs n Γ (a b : PTm (S n)) A B i :
Γ PBind PPi A B (PUniv i) -> Γ PBind PPi A B (PUniv i) ->
@ -92,20 +96,92 @@ with Eq : forall {n}, (fin n -> PTm n) -> PTm n -> PTm n -> PTm n -> Prop :=
Γ a b PBind PSig A B -> Γ a b PBind PSig A B ->
Γ PProj PR a PProj PR b subst_PTm (scons (PProj PL a) VarPTm) 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 -> Γ a b A ->
Γ B PUniv i -> Γ A B ->
Γ A B PUniv i ->
Γ a b B Γ a b B
| E_Bind_Proj1 n Γ p (A0 A1 : PTm n) B0 B1 i : (* Beta *)
Γ PBind p A0 B0 PBind p A1 B1 PUniv i -> | E_AppAbs n Γ (a : PTm (S n)) b A B i:
Γ A0 A1 PUniv 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_Bind_Proj2 n Γ p (a0 a1 A0 A1 : PTm n) B0 B1 i : | E_ProjPair1 n Γ (a b : PTm n) A B i :
Γ PBind p A0 B0 PBind p A1 B1 PUniv 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 ->
Γ PUniv i : PTm n PUniv j
| 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 i :
Γ ->
Γ A1 PUniv i ->
Γ A0 A1 ->
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
| 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 -> Γ 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 := with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
| Wff_Nil : | Wff_Nil :
@ -117,10 +193,18 @@ with Wff : forall {n}, (fin n -> PTm n) -> Prop :=
funcomp (ren_PTm shift) (scons A Γ) funcomp (ren_PTm shift) (scons A Γ)
where 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 Scheme wf_ind := Induction for Wff Sort Prop
with wt_ind := Induction for Wt 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 -> ...) /\ *)
(* (forall n Γ (A B : PTm n), Γ ⊢ A ≲ B -> ...). *)
(* Proof. apply wt_mutual. ... *)