Make the internal language even smaller
This commit is contained in:
parent
f402f4e528
commit
3b8fe388dc
4 changed files with 133 additions and 504 deletions
|
@ -19,29 +19,13 @@ Proof.
|
|||
exact (eq_refl).
|
||||
Qed.
|
||||
|
||||
Inductive TTag : Type :=
|
||||
| TPi : TTag
|
||||
| TSig : TTag.
|
||||
|
||||
Lemma congr_TPi : TPi = TPi.
|
||||
Proof.
|
||||
exact (eq_refl).
|
||||
Qed.
|
||||
|
||||
Lemma congr_TSig : TSig = TSig.
|
||||
Proof.
|
||||
exact (eq_refl).
|
||||
Qed.
|
||||
|
||||
Inductive PTm : Type :=
|
||||
| VarPTm : nat -> PTm
|
||||
| PAbs : PTm -> PTm
|
||||
| PApp : PTm -> PTm -> PTm
|
||||
| PPair : PTm -> PTm -> PTm
|
||||
| PProj : PTag -> PTm -> PTm
|
||||
| PConst : TTag -> PTm
|
||||
| PUniv : nat -> PTm
|
||||
| PBot : PTm.
|
||||
| PConst : nat -> PTm.
|
||||
|
||||
Lemma congr_PAbs {s0 : PTm} {t0 : PTm} (H0 : s0 = t0) : PAbs s0 = PAbs t0.
|
||||
Proof.
|
||||
|
@ -69,22 +53,12 @@ exact (eq_trans (eq_trans eq_refl (ap (fun x => PProj x s1) H0))
|
|||
(ap (fun x => PProj t0 x) H1)).
|
||||
Qed.
|
||||
|
||||
Lemma congr_PConst {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
|
||||
Lemma congr_PConst {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
|
||||
PConst s0 = PConst t0.
|
||||
Proof.
|
||||
exact (eq_trans eq_refl (ap (fun x => PConst x) H0)).
|
||||
Qed.
|
||||
|
||||
Lemma congr_PUniv {s0 : nat} {t0 : nat} (H0 : s0 = t0) : PUniv s0 = PUniv t0.
|
||||
Proof.
|
||||
exact (eq_trans eq_refl (ap (fun x => PUniv x) H0)).
|
||||
Qed.
|
||||
|
||||
Lemma congr_PBot : PBot = PBot.
|
||||
Proof.
|
||||
exact (eq_refl).
|
||||
Qed.
|
||||
|
||||
Lemma upRen_PTm_PTm (xi : nat -> nat) : nat -> nat.
|
||||
Proof.
|
||||
exact (up_ren xi).
|
||||
|
@ -98,8 +72,6 @@ Fixpoint ren_PTm (xi_PTm : nat -> nat) (s : PTm) {struct s} : PTm :=
|
|||
| PPair s0 s1 => PPair (ren_PTm xi_PTm s0) (ren_PTm xi_PTm s1)
|
||||
| PProj s0 s1 => PProj s0 (ren_PTm xi_PTm s1)
|
||||
| PConst s0 => PConst s0
|
||||
| PUniv s0 => PUniv s0
|
||||
| PBot => PBot
|
||||
end.
|
||||
|
||||
Lemma up_PTm_PTm (sigma : nat -> PTm) : nat -> PTm.
|
||||
|
@ -115,8 +87,6 @@ Fixpoint subst_PTm (sigma_PTm : nat -> PTm) (s : PTm) {struct s} : PTm :=
|
|||
| PPair s0 s1 => PPair (subst_PTm sigma_PTm s0) (subst_PTm sigma_PTm s1)
|
||||
| PProj s0 s1 => PProj s0 (subst_PTm sigma_PTm s1)
|
||||
| PConst s0 => PConst s0
|
||||
| PUniv s0 => PUniv s0
|
||||
| PBot => PBot
|
||||
end.
|
||||
|
||||
Lemma upId_PTm_PTm (sigma : nat -> PTm) (Eq : forall x, sigma x = VarPTm x) :
|
||||
|
@ -145,8 +115,6 @@ subst_PTm sigma_PTm s = s :=
|
|||
(idSubst_PTm sigma_PTm Eq_PTm s1)
|
||||
| PProj s0 s1 => congr_PProj (eq_refl s0) (idSubst_PTm sigma_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma upExtRen_PTm_PTm (xi : nat -> nat) (zeta : nat -> nat)
|
||||
|
@ -177,8 +145,6 @@ ren_PTm xi_PTm s = ren_PTm zeta_PTm s :=
|
|||
| PProj s0 s1 =>
|
||||
congr_PProj (eq_refl s0) (extRen_PTm xi_PTm zeta_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma upExt_PTm_PTm (sigma : nat -> PTm) (tau : nat -> PTm)
|
||||
|
@ -210,8 +176,6 @@ subst_PTm sigma_PTm s = subst_PTm tau_PTm s :=
|
|||
| PProj s0 s1 =>
|
||||
congr_PProj (eq_refl s0) (ext_PTm sigma_PTm tau_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma up_ren_ren_PTm_PTm (xi : nat -> nat) (zeta : nat -> nat)
|
||||
|
@ -242,8 +206,6 @@ Fixpoint compRenRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat)
|
|||
congr_PProj (eq_refl s0)
|
||||
(compRenRen_PTm xi_PTm zeta_PTm rho_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma up_ren_subst_PTm_PTm (xi : nat -> nat) (tau : nat -> PTm)
|
||||
|
@ -278,8 +240,6 @@ Fixpoint compRenSubst_PTm (xi_PTm : nat -> nat) (tau_PTm : nat -> PTm)
|
|||
congr_PProj (eq_refl s0)
|
||||
(compRenSubst_PTm xi_PTm tau_PTm theta_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma up_subst_ren_PTm_PTm (sigma : nat -> PTm) (zeta_PTm : nat -> nat)
|
||||
|
@ -325,8 +285,6 @@ ren_PTm zeta_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
|
|||
congr_PProj (eq_refl s0)
|
||||
(compSubstRen_PTm sigma_PTm zeta_PTm theta_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma up_subst_subst_PTm_PTm (sigma : nat -> PTm) (tau_PTm : nat -> PTm)
|
||||
|
@ -373,8 +331,6 @@ subst_PTm tau_PTm (subst_PTm sigma_PTm s) = subst_PTm theta_PTm s :=
|
|||
congr_PProj (eq_refl s0)
|
||||
(compSubstSubst_PTm sigma_PTm tau_PTm theta_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma renRen_PTm (xi_PTm : nat -> nat) (zeta_PTm : nat -> nat) (s : PTm) :
|
||||
|
@ -464,8 +420,6 @@ Fixpoint rinst_inst_PTm (xi_PTm : nat -> nat) (sigma_PTm : nat -> PTm)
|
|||
| PProj s0 s1 =>
|
||||
congr_PProj (eq_refl s0) (rinst_inst_PTm xi_PTm sigma_PTm Eq_PTm s1)
|
||||
| PConst s0 => congr_PConst (eq_refl s0)
|
||||
| PUniv s0 => congr_PUniv (eq_refl s0)
|
||||
| PBot => congr_PBot
|
||||
end.
|
||||
|
||||
Lemma rinstInst'_PTm (xi_PTm : nat -> nat) (s : PTm) :
|
||||
|
@ -527,6 +481,20 @@ Proof.
|
|||
exact (fun x => eq_refl).
|
||||
Qed.
|
||||
|
||||
Inductive TTag : Type :=
|
||||
| TPi : TTag
|
||||
| TSig : TTag.
|
||||
|
||||
Lemma congr_TPi : TPi = TPi.
|
||||
Proof.
|
||||
exact (eq_refl).
|
||||
Qed.
|
||||
|
||||
Lemma congr_TSig : TSig = TSig.
|
||||
Proof.
|
||||
exact (eq_refl).
|
||||
Qed.
|
||||
|
||||
Inductive Tm : Type :=
|
||||
| VarTm : nat -> Tm
|
||||
| Abs : Tm -> Tm
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue