Compare commits

...
Sign in to create a new pull request.

21 commits

Author SHA1 Message Date
d9d96d2c8b Merge pull request 'church' (#1) from church into main
Reviewed-on: #1
2025-01-20 20:20:55 -05:00
36427daa61 Refactor the equational theory to use the compile function 2025-01-20 20:11:51 -05:00
7850314935 Add stub for par compile 2025-01-20 16:49:47 -05:00
9c9ce52b63 Recover the contra lemmas 2025-01-20 16:28:44 -05:00
d68df5d0bc Add compile 2025-01-20 15:33:46 -05:00
f3718707f2 Add constants to the reduction semantics 2025-01-20 13:56:40 -05:00
0f1e85c853 Prove injectivity of Pair and Lambdas 2025-01-10 13:31:58 -05:00
fd8335a803 Add injectivity for pairs 2025-01-10 12:39:47 -05:00
34a0c2856e Prove most of the confluence results for eta reduction 2025-01-09 20:21:38 -05:00
e75d7745fe Finish normalization 2025-01-09 16:17:38 -05:00
0d3b751a33 . 2025-01-09 15:16:05 -05:00
7021497615 Finish adequacy 2025-01-09 15:15:11 -05:00
Yiyun Liu
bf2a369824 Generalize the model to talk about termination 2025-01-09 00:35:46 -05:00
Yiyun Liu
ec03826083 Add beta without the junk rules 2025-01-08 19:47:54 -05:00
9ab338c9e1 Add wn 2025-01-08 15:31:40 -05:00
602fe929bc Add pars_var_inv 2025-01-05 00:21:19 -05:00
05d330254e Fix ERed 2025-01-04 23:59:21 -05:00
a6fd48c009 Finish prov_extract 2025-01-04 23:38:44 -05:00
919f1513ce Show that prov is preserved by beta red and eta exp 2025-01-04 23:02:12 -05:00
f0da5c97bb Show that rpar preserves prov 2025-01-04 20:38:04 -05:00
ee7be7584c Remove unnecessary usage of Equations 2025-01-04 16:56:21 -05:00
5 changed files with 1910 additions and 435 deletions

View file

@ -3,14 +3,15 @@ Tm(VarTm) : Type
PTag : Type PTag : Type
TTag : Type TTag : Type
TPi : TTag
TSig : TTag
PL : PTag PL : PTag
PR : PTag PR : PTag
TPi : TTag
TSig : TTag
Abs : (bind Tm in Tm) -> Tm Abs : (bind Tm in Tm) -> Tm
App : Tm -> Tm -> Tm App : Tm -> Tm -> Tm
Pair : Tm -> Tm -> Tm Pair : Tm -> Tm -> Tm
Proj : PTag -> Tm -> Tm Proj : PTag -> Tm -> Tm
TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm TBind : TTag -> Tm -> (bind Tm in Tm) -> Tm
Bot : Tm Const : TTag -> Tm
Univ : nat -> Tm Univ : nat -> Tm
Bot : Tm

View file

@ -40,8 +40,9 @@ Inductive Tm (n_Tm : nat) : Type :=
| Pair : Tm n_Tm -> Tm n_Tm -> Tm n_Tm | Pair : Tm n_Tm -> Tm n_Tm -> Tm n_Tm
| Proj : PTag -> Tm n_Tm -> Tm n_Tm | Proj : PTag -> Tm n_Tm -> Tm n_Tm
| TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm | TBind : TTag -> Tm n_Tm -> Tm (S n_Tm) -> Tm n_Tm
| Bot : Tm n_Tm | Const : TTag -> Tm n_Tm
| Univ : nat -> Tm n_Tm. | Univ : nat -> Tm n_Tm
| Bot : Tm n_Tm.
Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)} Lemma congr_Abs {m_Tm : nat} {s0 : Tm (S m_Tm)} {t0 : Tm (S m_Tm)}
(H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0. (H0 : s0 = t0) : Abs m_Tm s0 = Abs m_Tm t0.
@ -83,9 +84,10 @@ exact (eq_trans
(ap (fun x => TBind m_Tm t0 t1 x) H2)). (ap (fun x => TBind m_Tm t0 t1 x) H2)).
Qed. Qed.
Lemma congr_Bot {m_Tm : nat} : Bot m_Tm = Bot m_Tm. Lemma congr_Const {m_Tm : nat} {s0 : TTag} {t0 : TTag} (H0 : s0 = t0) :
Const m_Tm s0 = Const m_Tm t0.
Proof. Proof.
exact (eq_refl). exact (eq_trans eq_refl (ap (fun x => Const m_Tm x) H0)).
Qed. Qed.
Lemma congr_Univ {m_Tm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) : Lemma congr_Univ {m_Tm : nat} {s0 : nat} {t0 : nat} (H0 : s0 = t0) :
@ -94,6 +96,11 @@ Proof.
exact (eq_trans eq_refl (ap (fun x => Univ m_Tm x) H0)). exact (eq_trans eq_refl (ap (fun x => Univ m_Tm x) H0)).
Qed. Qed.
Lemma congr_Bot {m_Tm : nat} : Bot m_Tm = Bot m_Tm.
Proof.
exact (eq_refl).
Qed.
Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) : Lemma upRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) :
fin (S m) -> fin (S n). fin (S m) -> fin (S n).
Proof. Proof.
@ -116,8 +123,9 @@ Fixpoint ren_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
| Proj _ s0 s1 => Proj n_Tm s0 (ren_Tm xi_Tm s1) | Proj _ s0 s1 => Proj n_Tm s0 (ren_Tm xi_Tm s1)
| TBind _ s0 s1 s2 => | TBind _ s0 s1 s2 =>
TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2) TBind n_Tm s0 (ren_Tm xi_Tm s1) (ren_Tm (upRen_Tm_Tm xi_Tm) s2)
| Bot _ => Bot n_Tm | Const _ s0 => Const n_Tm s0
| Univ _ s0 => Univ n_Tm s0 | Univ _ s0 => Univ n_Tm s0
| Bot _ => Bot n_Tm
end. end.
Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) : Lemma up_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) :
@ -143,8 +151,9 @@ Fixpoint subst_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
| Proj _ s0 s1 => Proj n_Tm s0 (subst_Tm sigma_Tm s1) | Proj _ s0 s1 => Proj n_Tm s0 (subst_Tm sigma_Tm s1)
| TBind _ s0 s1 s2 => | TBind _ s0 s1 s2 =>
TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2) TBind n_Tm s0 (subst_Tm sigma_Tm s1) (subst_Tm (up_Tm_Tm sigma_Tm) s2)
| Bot _ => Bot n_Tm | Const _ s0 => Const n_Tm s0
| Univ _ s0 => Univ n_Tm s0 | Univ _ s0 => Univ n_Tm s0
| Bot _ => Bot n_Tm
end. end.
Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm) Lemma upId_Tm_Tm {m_Tm : nat} (sigma : fin m_Tm -> Tm m_Tm)
@ -183,8 +192,9 @@ subst_Tm sigma_Tm s = s :=
| TBind _ s0 s1 s2 => | TBind _ s0 s1 s2 =>
congr_TBind (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1) congr_TBind (eq_refl s0) (idSubst_Tm sigma_Tm Eq_Tm s1)
(idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2) (idSubst_Tm (up_Tm_Tm sigma_Tm) (upId_Tm_Tm _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n) Lemma upExtRen_Tm_Tm {m : nat} {n : nat} (xi : fin m -> fin n)
@ -227,8 +237,9 @@ Fixpoint extRen_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
congr_TBind (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1) congr_TBind (eq_refl s0) (extRen_Tm xi_Tm zeta_Tm Eq_Tm s1)
(extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm) (extRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
(upExtRen_Tm_Tm _ _ Eq_Tm) s2) (upExtRen_Tm_Tm _ _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm) Lemma upExt_Tm_Tm {m : nat} {n_Tm : nat} (sigma : fin m -> Tm n_Tm)
@ -272,8 +283,9 @@ Fixpoint ext_Tm {m_Tm : nat} {n_Tm : nat} (sigma_Tm : fin m_Tm -> Tm n_Tm)
congr_TBind (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1) congr_TBind (eq_refl s0) (ext_Tm sigma_Tm tau_Tm Eq_Tm s1)
(ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm) (ext_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (upExt_Tm_Tm _ _ Eq_Tm)
s2) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l) Lemma up_ren_ren_Tm_Tm {k : nat} {l : nat} {m : nat} (xi : fin k -> fin l)
@ -317,8 +329,9 @@ Fixpoint compRenRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
congr_TBind (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1) congr_TBind (eq_refl s0) (compRenRen_Tm xi_Tm zeta_Tm rho_Tm Eq_Tm s1)
(compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm) (compRenRen_Tm (upRen_Tm_Tm xi_Tm) (upRen_Tm_Tm zeta_Tm)
(upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2) (upRen_Tm_Tm rho_Tm) (up_ren_ren _ _ _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat} Lemma up_ren_subst_Tm_Tm {k : nat} {l : nat} {m_Tm : nat}
@ -373,8 +386,9 @@ Fixpoint compRenSubst_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
(compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1) (compRenSubst_Tm xi_Tm tau_Tm theta_Tm Eq_Tm s1)
(compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm) (compRenSubst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm tau_Tm)
(up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2) (up_Tm_Tm theta_Tm) (up_ren_subst_Tm_Tm _ _ _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat} Lemma up_subst_ren_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@ -450,8 +464,9 @@ ren_Tm zeta_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
(compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1) (compSubstRen_Tm sigma_Tm zeta_Tm theta_Tm Eq_Tm s1)
(compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm) (compSubstRen_Tm (up_Tm_Tm sigma_Tm) (upRen_Tm_Tm zeta_Tm)
(up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2) (up_Tm_Tm theta_Tm) (up_subst_ren_Tm_Tm _ _ _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat} Lemma up_subst_subst_Tm_Tm {k : nat} {l_Tm : nat} {m_Tm : nat}
@ -528,8 +543,9 @@ subst_Tm tau_Tm (subst_Tm sigma_Tm s) = subst_Tm theta_Tm s :=
(compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1) (compSubstSubst_Tm sigma_Tm tau_Tm theta_Tm Eq_Tm s1)
(compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm) (compSubstSubst_Tm (up_Tm_Tm sigma_Tm) (up_Tm_Tm tau_Tm)
(up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2) (up_Tm_Tm theta_Tm) (up_subst_subst_Tm_Tm _ _ _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat} Lemma renRen_Tm {k_Tm : nat} {l_Tm : nat} {m_Tm : nat}
@ -644,8 +660,9 @@ Fixpoint rinst_inst_Tm {m_Tm : nat} {n_Tm : nat}
congr_TBind (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1) congr_TBind (eq_refl s0) (rinst_inst_Tm xi_Tm sigma_Tm Eq_Tm s1)
(rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm) (rinst_inst_Tm (upRen_Tm_Tm xi_Tm) (up_Tm_Tm sigma_Tm)
(rinstInst_up_Tm_Tm _ _ Eq_Tm) s2) (rinstInst_up_Tm_Tm _ _ Eq_Tm) s2)
| Bot _ => congr_Bot | Const _ s0 => congr_Const (eq_refl s0)
| Univ _ s0 => congr_Univ (eq_refl s0) | Univ _ s0 => congr_Univ (eq_refl s0)
| Bot _ => congr_Bot
end. end.
Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm) Lemma rinstInst'_Tm {m_Tm : nat} {n_Tm : nat} (xi_Tm : fin m_Tm -> fin n_Tm)
@ -844,9 +861,11 @@ Core.
Arguments VarTm {n_Tm}. Arguments VarTm {n_Tm}.
Arguments Bot {n_Tm}.
Arguments Univ {n_Tm}. Arguments Univ {n_Tm}.
Arguments Bot {n_Tm}. Arguments Const {n_Tm}.
Arguments TBind {n_Tm}. Arguments TBind {n_Tm}.

181
theories/compile.v Normal file
View file

@ -0,0 +1,181 @@
Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red.
Require Import ssreflect ssrbool.
From Hammer Require Import Tactics.
From stdpp Require Import relations (rtc(..)).
Module Compile.
Fixpoint F {n} (a : Tm n) : Tm n :=
match a with
| TBind p A B => Pair (Pair (Const p) (F A)) (Abs (F B))
| Const k => Const k
| Univ i => Univ i
| Abs a => Abs (F a)
| App a b => App (F a) (F b)
| VarTm i => VarTm i
| Pair a b => Pair (F a) (F b)
| Proj t a => Proj t (F a)
| Bot => Bot
end.
Lemma renaming n m (a : Tm n) (ξ : fin n -> fin m) :
F (ren_Tm ξ a)= ren_Tm ξ (F a).
Proof. move : m ξ. elim : n / a => //=; scongruence. Qed.
#[local]Hint Rewrite Compile.renaming : compile.
Lemma morphing n m (a : Tm n) (ρ0 ρ1 : fin n -> Tm m) :
(forall i, ρ0 i = F (ρ1 i)) ->
subst_Tm ρ0 (F a) = F (subst_Tm ρ1 a).
Proof.
move : m ρ0 ρ1. elim : n / a => n//=.
- hauto lq:on inv:option rew:db:compile unfold:funcomp.
- hauto lq:on rew:off.
- hauto lq:on rew:off.
- hauto lq:on.
- hauto lq:on inv:option rew:db:compile unfold:funcomp.
Qed.
Lemma substing n b (a : Tm (S n)) :
subst_Tm (scons (F b) VarTm) (F a) = F (subst_Tm (scons b VarTm) a).
Proof.
apply morphing.
case => //=.
Qed.
End Compile.
#[export] Hint Rewrite Compile.renaming Compile.morphing : compile.
Module Join.
Definition R {n} (a b : Tm n) := join (Compile.F a) (Compile.F b).
Lemma BindInj n p0 p1 (A0 A1 : Tm n) B0 B1 :
R (TBind p0 A0 B0) (TBind p1 A1 B1) -> p0 = p1 /\ R A0 A1 /\ R B0 B1.
Proof.
rewrite /R /= !join_pair_inj.
move => [[/join_const_inj h0 h1] h2].
apply abs_eq in h2.
evar (t : Tm (S n)).
have : join (App (ren_Tm shift (Abs (Compile.F B1))) (VarTm var_zero)) t by
apply Join.FromPar; apply Par.AppAbs; auto using Par.refl.
subst t. rewrite -/ren_Tm.
move : h2. move /join_transitive => /[apply]. asimpl => h2.
tauto.
Qed.
Lemma UnivInj n i j : R (Univ i : Tm n) (Univ j) -> i = j.
Proof. hauto l:on use:join_univ_inj. Qed.
Lemma transitive n (a b c : Tm n) :
R a b -> R b c -> R a c.
Proof. hauto l:on use:join_transitive unfold:R. Qed.
Lemma symmetric n (a b : Tm n) :
R a b -> R b a.
Proof. hauto l:on use:join_symmetric. Qed.
Lemma reflexive n (a : Tm n) :
R a a.
Proof. hauto l:on use:join_refl. Qed.
Lemma substing n m (a b : Tm n) (ρ : fin n -> Tm m) :
R a b -> R (subst_Tm ρ a) (subst_Tm ρ b).
Proof.
rewrite /R.
rewrite /join.
move => [C [h0 h1]].
repeat (rewrite <- Compile.morphing with (ρ0 := funcomp Compile.F ρ); last by reflexivity).
hauto lq:on use:join_substing.
Qed.
End Join.
Module Equiv.
Inductive R {n} : Tm n -> Tm n -> Prop :=
(***************** Beta ***********************)
| AppAbs a b :
R (App (Abs a) b) (subst_Tm (scons b VarTm) a)
| ProjPair p a b :
R (Proj p (Pair a b)) (if p is PL then a else b)
(****************** Eta ***********************)
| AppEta a :
R a (Abs (App (ren_Tm shift a) (VarTm var_zero)))
| PairEta a :
R a (Pair (Proj PL a) (Proj PR a))
(*************** Congruence ********************)
| Var i : R (VarTm i) (VarTm i)
| AbsCong a b :
R a b ->
R (Abs a) (Abs b)
| AppCong a0 a1 b0 b1 :
R a0 a1 ->
R b0 b1 ->
R (App a0 b0) (App a1 b1)
| PairCong a0 a1 b0 b1 :
R a0 a1 ->
R b0 b1 ->
R (Pair a0 b0) (Pair a1 b1)
| ProjCong p a0 a1 :
R a0 a1 ->
R (Proj p a0) (Proj p a1)
| BindCong p A0 A1 B0 B1:
R A0 A1 ->
R B0 B1 ->
R (TBind p A0 B0) (TBind p A1 B1)
| UnivCong i :
R (Univ i) (Univ i).
End Equiv.
Module EquivJoin.
Lemma FromEquiv n (a b : Tm n) : Equiv.R a b -> Join.R a b.
Proof.
move => h. elim : n a b /h => n.
- move => a b.
rewrite /Join.R /join /=.
eexists. split. apply relations.rtc_once.
apply Par.AppAbs; auto using Par.refl.
rewrite Compile.substing.
apply relations.rtc_refl.
- move => p a b.
apply Join.FromPar.
simpl. apply : Par.ProjPair'; auto using Par.refl.
case : p => //=.
- move => a. apply Join.FromPar => /=.
apply : Par.AppEta'; auto using Par.refl.
by autorewrite with compile.
- move => a. apply Join.FromPar => /=.
apply : Par.PairEta; auto using Par.refl.
- hauto l:on use:Join.FromPar, Par.Var.
- hauto lq:on use:Join.AbsCong.
- qauto l:on use:Join.AppCong.
- qauto l:on use:Join.PairCong.
- qauto use:Join.ProjCong.
- rewrite /Join.R => p A0 A1 B0 B1 _ hA _ hB /=.
sfirstorder use:Join.PairCong,Join.AbsCong,Join.FromPar,Par.ConstCong.
- hauto l:on.
Qed.
End EquivJoin.
Lemma compile_rpar n (a b : Tm n) : RPar'.R a b -> RPar'.R (Compile.F a) (Compile.F b).
Proof.
move => h. elim : n a b /h.
- move => n a0 a1 b0 b1 ha iha hb ihb /=.
rewrite -Compile.substing.
apply RPar'.AppAbs => //.
- hauto q:on use:RPar'.ProjPair'.
- qauto ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
- hauto lq:on ctrs:RPar'.R.
Qed.
Lemma compile_rpars n (a b : Tm n) : rtc RPar'.R a b -> rtc RPar'.R (Compile.F a) (Compile.F b).
Proof. induction 1; hauto lq:on ctrs:rtc use:compile_rpar. Qed.

File diff suppressed because it is too large Load diff

View file

@ -1,23 +1,26 @@
Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax. Require Import Autosubst2.core Autosubst2.fintype Autosubst2.syntax fp_red compile.
Require Import 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.
Require Import Logic.PropExtensionality (propositional_extensionality). Require Import Logic.PropExtensionality (propositional_extensionality).
From stdpp Require Import relations (rtc(..), rtc_subrel). From stdpp Require Import relations (rtc(..), rtc_subrel).
Import Psatz. Import Psatz.
Definition ProdSpace (PA : Tm 0 -> Prop)
(PF : Tm 0 -> (Tm 0 -> Prop) -> Prop) b : Prop := Definition ProdSpace {n} (PA : Tm n -> Prop)
(PF : Tm n -> (Tm n -> Prop) -> Prop) b : Prop :=
forall a PB, PA a -> PF a PB -> PB (App b a). forall a PB, PA a -> PF a PB -> PB (App b a).
Definition SumSpace (PA : Tm 0 -> Prop) Definition SumSpace {n} (PA : Tm n -> Prop)
(PF : Tm 0 -> (Tm 0 -> Prop) -> Prop) t : Prop := (PF : Tm n -> (Tm n -> Prop) -> Prop) t : Prop :=
exists a b, rtc RPar.R t (Pair a b) /\ PA a /\ (forall PB, PF a PB -> PB b). wne t \/ exists a b, rtc RPar'.R t (Pair a b) /\ PA a /\ (forall PB, PF a PB -> PB b).
Definition BindSpace p := if p is TPi then ProdSpace else SumSpace. Definition BindSpace {n} p := if p is TPi then @ProdSpace n else SumSpace.
Reserved Notation "⟦ A ⟧ i ;; I ↘ S" (at level 70). Reserved Notation "⟦ A ⟧ i ;; I ↘ S" (at level 70).
Inductive InterpExt i (I : nat -> Tm 0 -> Prop) : Tm 0 -> (Tm 0 -> Prop) -> Prop := Inductive InterpExt {n} i (I : nat -> Tm n -> Prop) : Tm n -> (Tm n -> Prop) -> Prop :=
| InterpExt_Ne A :
ne A ->
A i ;; I wne
| 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) ->
@ -29,12 +32,12 @@ Inductive InterpExt i (I : nat -> Tm 0 -> Prop) : Tm 0 -> (Tm 0 -> Prop) -> Prop
Univ j i ;; I (I j) Univ j i ;; I (I j)
| InterpExt_Step A A0 PA : | InterpExt_Step A A0 PA :
RPar.R A A0 -> RPar'.R A A0 ->
A0 i ;; I PA -> A0 i ;; I PA ->
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' i I j (PF : Tm 0 -> Prop) : Lemma InterpExt_Univ' n i I j (PF : Tm n -> Prop) :
PF = I j -> PF = I j ->
j < i -> j < i ->
Univ j i ;; I PF. Univ j i ;; I PF.
@ -42,28 +45,29 @@ Proof. hauto lq:on ctrs:InterpExt. Qed.
Infix "<?" := Compare_dec.lt_dec (at level 60). Infix "<?" := Compare_dec.lt_dec (at level 60).
Equations InterpUnivN (i : nat) : Tm 0 -> (Tm 0 -> Prop) -> Prop by wf i lt := Equations InterpUnivN n (i : nat) : Tm n -> (Tm n -> Prop) -> Prop by wf i lt :=
InterpUnivN i := @InterpExt i InterpUnivN n i := @InterpExt n i
(fun j A => (fun j A =>
match j <? i with match j <? i with
| left _ => exists PA, InterpUnivN j A PA | left _ => exists PA, InterpUnivN n j A PA
| right _ => False | right _ => False
end). end).
Arguments InterpUnivN . Arguments InterpUnivN {n}.
Lemma InterpExt_lt_impl i I I' A (PA : Tm 0 -> Prop) : Lemma InterpExt_lt_impl n i I I' A (PA : Tm n -> Prop) :
(forall j, j < i -> I j = I' j) -> (forall j, j < i -> I j = I' j) ->
A i ;; I PA -> A i ;; I PA ->
A i ;; I' PA. A i ;; I' PA.
Proof. Proof.
move => hI h. move => hI h.
elim : A PA /h. elim : A PA /h.
- hauto q:on ctrs:InterpExt.
- hauto lq:on rew:off ctrs:InterpExt. - hauto lq:on rew:off ctrs:InterpExt.
- hauto q:on ctrs:InterpExt. - hauto q:on ctrs:InterpExt.
- hauto lq:on ctrs:InterpExt. - hauto lq:on ctrs:InterpExt.
Qed. Qed.
Lemma InterpExt_lt_eq i I I' A (PA : Tm 0 -> Prop) : Lemma InterpExt_lt_eq n i I I' A (PA : Tm n -> Prop) :
(forall j, j < i -> I j = I' j) -> (forall j, j < i -> I j = I' j) ->
A i ;; I PA = A i ;; I PA =
A i ;; I' PA. A i ;; I' PA.
@ -75,8 +79,8 @@ Qed.
Notation "⟦ A ⟧ i ↘ S" := (InterpUnivN i A S) (at level 70). Notation "⟦ A ⟧ i ↘ S" := (InterpUnivN i A S) (at level 70).
Lemma InterpUnivN_nolt i : Lemma InterpUnivN_nolt n i :
InterpUnivN i = InterpExt i (fun j (A : Tm 0) => exists PA, A j PA). @InterpUnivN n i = @InterpExt n i (fun j (A : Tm n) => exists PA, A j PA).
Proof. Proof.
simp InterpUnivN. simp InterpUnivN.
extensionality A. extensionality PA. extensionality A. extensionality PA.
@ -89,12 +93,12 @@ Qed.
#[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv. #[export]Hint Rewrite @InterpUnivN_nolt : InterpUniv.
Lemma RPar_substone n (a b : Tm (S n)) (c : Tm n): Lemma RPar_substone n (a b : Tm (S n)) (c : Tm n):
RPar.R a b -> RPar.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b). RPar'.R a b -> RPar'.R (subst_Tm (scons c VarTm) a) (subst_Tm (scons c VarTm) b).
Proof. hauto l:on inv:option use:RPar.substing, RPar.refl. Qed. Proof. hauto l:on inv:option use:RPar'.substing, RPar'.refl. Qed.
Lemma InterpExt_Bind_inv p i I (A : Tm 0) B P Lemma InterpExt_Bind_inv n p i I (A : Tm n) B P
(h : TBind p A B i ;; I P) : (h : TBind p A B i ;; I P) :
exists (PA : Tm 0 -> Prop) (PF : Tm 0 -> (Tm 0 -> Prop) -> Prop), exists (PA : Tm n -> Prop) (PF : Tm n -> (Tm n -> Prop) -> Prop),
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) /\
(forall a PB, PF a PB -> subst_Tm (scons a VarTm) B i ;; I PB) /\ (forall a PB, PF a PB -> subst_Tm (scons a VarTm) B i ;; I PB) /\
@ -103,24 +107,35 @@ Proof.
move E : (TBind p A B) h => T h. move E : (TBind p A B) h => T h.
move : A B E. move : A B E.
elim : T P / h => //. elim : T P / h => //.
- move => //= *. scongruence.
- hauto l:on. - hauto l:on.
- move => A A0 PA hA hA0 hPi A1 B ?. subst. - move => A A0 PA hA hA0 hPi A1 B ?. subst.
elim /RPar.inv : hA => //= _ p0 A2 A3 B0 B1 hA1 hB0 [*]. subst. elim /RPar'.inv : hA => //= _ p0 A2 A3 B0 B1 hA1 hB0 [*]. subst.
hauto lq:on ctrs:InterpExt use:RPar_substone. hauto lq:on ctrs:InterpExt use:RPar_substone.
Qed. Qed.
Lemma InterpExt_Univ_inv i I j P Lemma InterpExt_Ne_inv n i A I P
(h : Univ j i ;; I P) : (h : A : Tm n i ;; I P) :
ne A ->
P = wne.
Proof.
elim : A P / h => //=.
qauto l:on ctrs:prov inv:prov use:nf_refl.
Qed.
Lemma InterpExt_Univ_inv n i I j P
(h : Univ j : Tm n i ;; I P) :
P = I j /\ j < i. P = I j /\ j < i.
Proof. Proof.
move : h. move : h.
move E : (Univ j) => T h. move : j E. move E : (Univ j) => T h. move : j E.
elim : T P /h => //. elim : T P /h => //.
- move => //= *. scongruence.
- hauto l:on. - hauto l:on.
- hauto lq:on rew:off inv:RPar.R. - hauto lq:on rew:off inv:RPar'.R.
Qed. Qed.
Lemma InterpExt_Bind_nopf p i I (A : Tm 0) B PA : Lemma InterpExt_Bind_nopf n p i I (A : Tm n) B PA :
A i ;; I PA -> A i ;; I PA ->
(forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i ;; I PB) -> (forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i ;; I PB) ->
TBind p A B i ;; I (BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i ;; I PB)). TBind p A B i ;; I (BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i ;; I PB)).
@ -128,7 +143,7 @@ Proof.
move => h0 h1. apply InterpExt_Bind =>//. move => h0 h1. apply InterpExt_Bind =>//.
Qed. Qed.
Lemma InterpUnivN_Fun_nopf p i (A : Tm 0) B PA : Lemma InterpUnivN_Fun_nopf n p i (A : Tm n) B PA :
A i PA -> A i PA ->
(forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i PB) -> (forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i PB) ->
TBind p A B i (BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i PB)). TBind p A B i (BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i PB)).
@ -136,7 +151,7 @@ Proof.
hauto l:on use:InterpExt_Bind_nopf rew:db:InterpUniv. hauto l:on use:InterpExt_Bind_nopf rew:db:InterpUniv.
Qed. Qed.
Lemma InterpExt_cumulative i j I (A : Tm 0) PA : Lemma InterpExt_cumulative n i j I (A : Tm n) PA :
i <= j -> i <= j ->
A i ;; I PA -> A i ;; I PA ->
A j ;; I PA. A j ;; I PA.
@ -146,61 +161,87 @@ Proof.
hauto l:on ctrs:InterpExt solve+:(by lia). hauto l:on ctrs:InterpExt solve+:(by lia).
Qed. Qed.
Lemma InterpUnivN_cumulative i (A : Tm 0) PA : Lemma InterpUnivN_cumulative n i (A : Tm n) PA :
A i PA -> forall j, i <= j -> A i PA -> forall j, i <= j ->
A j PA. A j PA.
Proof. Proof.
hauto l:on rew:db:InterpUniv use:InterpExt_cumulative. hauto l:on rew:db:InterpUniv use:InterpExt_cumulative.
Qed. Qed.
Lemma InterpExt_preservation i I (A : Tm 0) B P (h : InterpExt i I A P) : Lemma InterpExt_preservation n i I (A : Tm n) B P (h : InterpExt i I A P) :
RPar.R A B -> RPar'.R A B ->
B i ;; I P. B i ;; I P.
Proof. Proof.
move : B. move : B.
elim : A P / h; auto. elim : A P / h; auto.
- hauto lq:on use:nf_refl ctrs:InterpExt.
- move => p A B PA PF hPA ihPA hPB hPB' ihPB T hT. - move => p A B PA PF hPA ihPA hPB hPB' ihPB T hT.
elim /RPar.inv : hT => //. elim /RPar'.inv : hT => //.
move => hPar p0 A0 A1 B0 B1 h0 h1 [? ?] ? ?; subst. move => hPar p0 A0 A1 B0 B1 h0 h1 [? ?] ? ?; subst.
apply InterpExt_Bind; auto => a PB hPB0. apply InterpExt_Bind; auto => a PB hPB0.
apply : ihPB; eauto. apply : ihPB; eauto.
sfirstorder use:RPar.cong, RPar.refl. sfirstorder use:RPar'.cong, RPar'.refl.
- hauto lq:on inv:RPar.R ctrs:InterpExt. - hauto lq:on inv:RPar'.R ctrs:InterpExt.
- move => A B P h0 h1 ih1 C hC. - move => A B P h0 h1 ih1 C hC.
have [D [h2 h3]] := RPar_diamond _ _ _ _ h0 hC. have [D [h2 h3]] := RPar'_diamond _ _ _ _ h0 hC.
hauto lq:on ctrs:InterpExt. hauto lq:on ctrs:InterpExt.
Qed. Qed.
Lemma InterpUnivN_preservation i (A : Tm 0) B P (h : A i P) : Lemma InterpUnivN_preservation n i (A : Tm n) B P (h : A i P) :
RPar.R A B -> RPar'.R A B ->
B i P. B i P.
Proof. hauto l:on rew:db:InterpUnivN use: InterpExt_preservation. Qed. Proof. hauto l:on rew:db:InterpUnivN use: InterpExt_preservation. Qed.
Lemma InterpExt_back_preservation_star i I (A : Tm 0) B P (h : B i ;; I P) : Lemma InterpExt_back_preservation_star n i I (A : Tm n) B P (h : B i ;; I P) :
rtc RPar.R A B -> rtc RPar'.R A B ->
A i ;; I P. A i ;; I P.
Proof. induction 1; hauto l:on ctrs:InterpExt. Qed. Proof. induction 1; hauto l:on ctrs:InterpExt. Qed.
Lemma InterpExt_preservation_star i I (A : Tm 0) B P (h : A i ;; I P) : Lemma InterpExt_preservation_star n i I (A : Tm n) B P (h : A i ;; I P) :
rtc RPar.R A B -> rtc RPar'.R A B ->
B i ;; I P. B i ;; I P.
Proof. induction 1; hauto l:on use:InterpExt_preservation. Qed. Proof. induction 1; hauto l:on use:InterpExt_preservation. Qed.
Lemma InterpUnivN_preservation_star i (A : Tm 0) B P (h : A i P) : Lemma InterpUnivN_preservation_star n i (A : Tm n) B P (h : A i P) :
rtc RPar.R A B -> rtc RPar'.R A B ->
B i P. B i P.
Proof. hauto l:on rew:db:InterpUnivN use:InterpExt_preservation_star. Qed. Proof. hauto l:on rew:db:InterpUnivN use:InterpExt_preservation_star. Qed.
Lemma InterpUnivN_back_preservation_star i (A : Tm 0) B P (h : B i P) : Lemma InterpUnivN_back_preservation_star n i (A : Tm n) B P (h : B i P) :
rtc RPar.R A B -> rtc RPar'.R A B ->
A i P. A i P.
Proof. hauto l:on rew:db:InterpUnivN use:InterpExt_back_preservation_star. Qed. Proof. hauto l:on rew:db:InterpUnivN use:InterpExt_back_preservation_star. Qed.
Lemma InterpExtInv i I (A : Tm 0) PA : Function hfb {n} (A : Tm n) :=
match A with
| TBind _ _ _ => true
| Univ _ => true
| _ => ne A
end.
Inductive hfb_case {n} : Tm n -> Prop :=
| hfb_bind p A B :
hfb_case (TBind p A B)
| hfb_univ i :
hfb_case (Univ i)
| hfb_ne A :
ne A ->
hfb_case A.
Derive Dependent Inversion hfb_inv with (forall n (a : Tm n), hfb_case a) Sort Prop.
Lemma ne_hfb {n} (A : Tm n) : ne A -> hfb A.
Proof. case : A => //=. Qed.
Lemma hfb_caseP {n} (A : Tm n) : hfb A -> hfb_case A.
Proof. hauto lq:on ctrs:hfb_case inv:Tm use:ne_hfb. Qed.
Lemma InterpExtInv n i I (A : Tm n) PA :
A i ;; I PA -> A i ;; I PA ->
exists B, hfb B /\ rtc RPar.R A B /\ B i ;; I PA. exists B, hfb B /\ rtc RPar'.R A B /\ B i ;; I PA.
Proof. Proof.
move => h. elim : A PA /h. move => h. elim : A PA /h.
- hauto q:on ctrs:InterpExt, rtc use:ne_hfb.
- move => p A B PA PF hPA _ hPF hPF0 _. - move => p A B PA PF hPA _ hPF hPF0 _.
exists (TBind p A B). repeat split => //=. exists (TBind p A B). repeat split => //=.
apply rtc_refl. apply rtc_refl.
@ -210,17 +251,26 @@ Proof.
- hauto lq:on ctrs:rtc. - hauto lq:on ctrs:rtc.
Qed. Qed.
Lemma RPars_Pars (A B : Tm 0) : Lemma RPar'_Par n (A B : Tm n) :
rtc RPar.R A B -> RPar'.R A B ->
Par.R A B.
Proof. induction 1; hauto lq:on ctrs:Par.R. Qed.
Lemma RPar's_Pars n (A B : Tm n) :
rtc RPar'.R A B ->
rtc Par.R A B. rtc Par.R A B.
Proof. hauto lq:on use:RPar_Par, rtc_subrel. Qed. Proof. hauto lq:on use:RPar'_Par, rtc_subrel. Qed.
Lemma RPars_join (A B : Tm 0) : Lemma RPar's_join n (A B : Tm n) :
rtc RPar.R A B -> join A B. rtc RPar'.R A B -> Join.R A B.
Proof. hauto lq:on ctrs:rtc use:RPars_Pars. Qed. Proof.
rewrite /Join.R => h.
have {}h : rtc RPar'.R (Compile.F A) (Compile.F B) by eauto using compile_rpars.
rewrite /join. eauto using RPar's_Pars, rtc_refl.
Qed.
Lemma bindspace_iff p (PA : Tm 0 -> Prop) PF PF0 b : Lemma bindspace_iff n p (PA : Tm n -> Prop) PF PF0 b :
(forall (a : Tm 0) (PB PB0 : Tm 0 -> Prop), PF a PB -> PF0 a PB0 -> PB = PB0) -> (forall (a : Tm n) (PB PB0 : Tm n -> Prop), PF a PB -> PF0 a PB0 -> PB = PB0) ->
(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, PA a -> exists PB0, PF0 a PB0) ->
(BindSpace p PA PF b <-> BindSpace p PA PF0 b). (BindSpace p PA PF b <-> BindSpace p PA PF0 b).
@ -241,23 +291,105 @@ Proof.
hauto lq:on rew:off. hauto lq:on rew:off.
Qed. Qed.
Lemma InterpExt_Join i I (A B : Tm 0) PA PB : Lemma ne_prov_inv n (a : Tm n) :
ne a -> (exists i, prov (VarTm i) a) \/ prov Bot a.
Proof.
elim : n /a => //=.
- hauto lq:on ctrs:prov.
- hauto lq:on rew:off ctrs:prov b:on.
- hauto lq:on ctrs:prov.
- move => n k.
have : @prov n Bot Bot by auto using P_Bot.
eauto.
Qed.
Lemma ne_pars_inv n (a b : Tm n) :
ne a -> rtc Par.R a b -> (exists i, prov (VarTm i) b) \/ prov Bot b.
Proof.
move /ne_prov_inv.
sfirstorder use:prov_pars.
Qed.
Lemma ne_pars_extract n (a b : Tm n) :
ne a -> rtc Par.R a b -> (exists i, extract b = (VarTm i)) \/ extract b = Bot.
Proof. hauto lq:on rew:off use:ne_pars_inv, prov_extract. Qed.
Lemma const_pars_extract n k b :
rtc Par.R (Const k : Tm n) b -> extract b = Const k.
Proof. hauto l:on use:pars_const_inv, prov_extract. Qed.
Lemma compile_ne n (a : Tm n) :
ne a = ne (Compile.F a) /\ nf a = nf (Compile.F a).
Proof.
elim : n / a => //=; sfirstorder b:on.
Qed.
Lemma join_univ_pi_contra n p (A : Tm n) B i :
Join.R (TBind p A B) (Univ i) -> False.
Proof.
rewrite /Join.R /=.
rewrite !pair_eq.
move => [[h _ ]_ ].
move : h => [C [h0 h1]].
have ? : extract C = Const p by hauto l:on use:pars_const_inv.
have h : prov (Univ i : Tm n) (Proj PL (Proj PL (Univ i))) by hauto lq:on ctrs:prov.
have {h} : prov (Univ i) C by eauto using prov_pars.
move /prov_extract=>/=. congruence.
Qed.
Lemma join_bind_ne_contra n p (A : Tm n) B C :
ne C ->
Join.R (TBind p A B) C -> False.
Proof.
rewrite /Join.R. move => hC /=.
rewrite !pair_eq.
move => [[[D [h0 h1]] _] _].
have {}hC : ne (Compile.F C) by hauto lq:on use:compile_ne.
have {}hC : ne (Proj PL (Proj PL (Compile.F C))) by scongruence.
have : (exists i, extract D = (VarTm i)) \/ extract D = Bot by eauto using ne_pars_extract.
have : extract D = Const p by eauto using const_pars_extract.
sfirstorder.
Qed.
Lemma join_univ_ne_contra n i C :
ne C ->
Join.R (Univ i : Tm n) C -> False.
Proof.
move => hC [D [h0 h1]].
move /pars_univ_inv : h0 => ?.
have : (exists i, extract D = (VarTm i)) \/ exists k, extract D =(Const k) by hauto q:on use:ne_pars_extract, compile_ne.
sfirstorder.
Qed.
#[export]Hint Resolve join_univ_ne_contra join_bind_ne_contra join_univ_pi_contra Join.symmetric Join.transitive : join.
Lemma InterpExt_Join n i I (A B : Tm n) PA PB :
A i ;; I PA -> A i ;; I PA ->
B i ;; I PB -> B i ;; I PB ->
join A B -> Join.R A B ->
PA = PB. PA = PB.
Proof. Proof.
move => h. move : B PB. elim : A PA /h. move => h. move : B PB. elim : A PA /h.
- move => A hA B PB /InterpExtInv.
move => [B0 []].
move /hfb_caseP. elim/hfb_inv => _.
+ move => p A0 B1 ? [/RPar's_join h0 h1] h2. subst. exfalso.
eauto with join.
+ move => ? ? [/RPar's_join *]. subst. exfalso.
eauto with join.
+ hauto lq:on use:InterpExt_Ne_inv.
- move => p A B PA PF hPA ihPA hTot hRes ihPF U PU /InterpExtInv. - move => p A B PA PF hPA ihPA hTot hRes ihPF U PU /InterpExtInv.
move => [B0 []]. move => [B0 []].
case : B0 => //=. move /hfb_caseP.
+ move => p0 A0 B0 _ [hr hPi]. elim /hfb_inv => _.
rename B0 into B00.
+ move => p0 A0 B0 ? [hr hPi]. subst.
move /InterpExt_Bind_inv : hPi. move /InterpExt_Bind_inv : hPi.
move => [PA0][PF0][hPA0][hTot0][hRes0]?. subst. move => [PA0][PF0][hPA0][hTot0][hRes0]?. subst.
move => hjoin. move => hjoin.
have{}hr : join U (TBind p0 A0 B0) by auto using RPars_join. have{}hr : Join.R U (TBind p0 A0 B0) by auto using RPar's_join.
have hj : join (TBind p A B) (TBind p0 A0 B0) by eauto using join_transitive. have hj : Join.R (TBind p A B) (TBind p0 A0 B0) by eauto using Join.transitive.
have {hj} : p0 = p /\ join A A0 /\ join B B0 by hauto l:on use:join_pi_inj. have {hj} : p0 = p /\ Join.R A A0 /\ Join.R B B0 by hauto l:on use:Join.BindInj.
move => [? [h0 h1]]. subst. move => [? [h0 h1]]. subst.
have ? : PA0 = PA by hauto l:on. subst. have ? : PA0 = PA by hauto l:on. subst.
rewrite /ProdSpace. rewrite /ProdSpace.
@ -266,66 +398,70 @@ Proof.
apply bindspace_iff; eauto. apply bindspace_iff; eauto.
move => a PB PB0 hPB hPB0. move => a PB PB0 hPB hPB0.
apply : ihPF; eauto. apply : ihPF; eauto.
by apply join_substing. rewrite /Join.R.
+ move => j _. rewrite -!Compile.substing.
hauto l:on use:join_substing.
+ move => j ?. subst.
move => [h0 h1] h. move => [h0 h1] h.
have ? : join U (Univ j) by eauto using RPars_join. have ? : Join.R U (Univ j) by eauto using RPar's_join.
have : join (TBind p A B) (Univ j) by eauto using join_transitive. have : Join.R (TBind p A B) (Univ j) by eauto using Join.transitive.
move => ?. exfalso. move => ?. exfalso.
eauto using join_univ_pi_contra. hauto l: on use: join_univ_pi_contra.
+ move => A0 ? ? [/RPar's_join ?]. subst.
move => _ ?. exfalso. eauto with join.
- move => j ? B PB /InterpExtInv. - move => j ? B PB /InterpExtInv.
move => [+ []]. case => //=. move => [? []]. move/hfb_caseP.
elim /hfb_inv => //= _.
+ move => p A0 B0 _ []. + move => p A0 B0 _ [].
move /RPars_join => *. move /RPar's_join => *.
have ? : join (TBind p A0 B0) (Univ j) by eauto using join_symmetric, join_transitive. exfalso. eauto with join.
exfalso. + move => m _ [/RPar's_join h0 + h1].
eauto using join_univ_pi_contra. have /join_univ_inj {h0 h1} ? : Join.R (Univ j : Tm n) (Univ m) by eauto using Join.transitive.
+ move => m _ [/RPars_join h0 + h1].
have /join_univ_inj {h0 h1} ? : join (Univ j : Tm 0) (Univ m) by eauto using join_transitive.
subst. subst.
move /InterpExt_Univ_inv. firstorder. move /InterpExt_Univ_inv. firstorder.
+ move => A ? ? [/RPar's_join] *. subst. exfalso. eauto with join.
- move => A A0 PA h. - move => A A0 PA h.
have /join_symmetric {}h : join A A0 by hauto lq:on ctrs:rtc use:RPar_Par, relations.rtc_once. have /Join.symmetric {}h : Join.R A A0 by hauto lq:on ctrs:rtc use:RPar's_join, relations.rtc_once.
eauto using join_transitive. eauto with join.
Qed. Qed.
Lemma InterpUniv_Join i (A B : Tm 0) PA PB : Lemma InterpUniv_Join n i (A B : Tm n) PA PB :
A i PA -> A i PA ->
B i PB -> B i PB ->
join A B -> Join.R A B ->
PA = PB. PA = PB.
Proof. hauto l:on use:InterpExt_Join rew:db:InterpUniv. Qed. Proof. hauto l:on use:InterpExt_Join rew:db:InterpUniv. Qed.
Lemma InterpUniv_Bind_inv p i (A : Tm 0) B P Lemma InterpUniv_Bind_inv n p i (A : Tm n) B P
(h : TBind p A B i P) : (h : TBind p A B i P) :
exists (PA : Tm 0 -> Prop) (PF : Tm 0 -> (Tm 0 -> Prop) -> Prop), exists (PA : Tm n -> Prop) (PF : Tm n -> (Tm n -> Prop) -> Prop),
A i PA /\ A i PA /\
(forall a, PA a -> exists PB, PF a PB) /\ (forall a, PA a -> exists PB, PF a PB) /\
(forall a PB, PF a PB -> subst_Tm (scons a VarTm) B i PB) /\ (forall a PB, PF a PB -> subst_Tm (scons a VarTm) B i PB) /\
P = BindSpace p PA PF. P = BindSpace p PA PF.
Proof. hauto l:on use:InterpExt_Bind_inv rew:db:InterpUniv. Qed. Proof. hauto l:on use:InterpExt_Bind_inv rew:db:InterpUniv. Qed.
Lemma InterpUniv_Univ_inv i j P Lemma InterpUniv_Univ_inv n i j P
(h : Univ j i P) : (h : Univ j i P) :
P = (fun (A : Tm 0) => exists PA, A j PA) /\ j < i. P = (fun (A : Tm n) => exists PA, A j PA) /\ j < i.
Proof. hauto l:on use:InterpExt_Univ_inv rew:db:InterpUniv. Qed. Proof. hauto l:on use:InterpExt_Univ_inv rew:db:InterpUniv. Qed.
Lemma InterpExt_Functional i I (A B : Tm 0) PA PB : Lemma InterpExt_Functional n i I (A B : Tm n) PA PB :
A i ;; I PA -> A i ;; I PA ->
A i ;; I PB -> A i ;; I PB ->
PA = PB. PA = PB.
Proof. hauto use:InterpExt_Join, join_refl. Qed. Proof. hauto use:InterpExt_Join, join_refl. Qed.
Lemma InterpUniv_Functional i (A : Tm 0) PA PB : Lemma InterpUniv_Functional n i (A : Tm n) PA PB :
A i PA -> A i PA ->
A i PB -> A i PB ->
PA = PB. PA = PB.
Proof. hauto use:InterpExt_Functional rew:db:InterpUniv. Qed. Proof. hauto use:InterpExt_Functional rew:db:InterpUniv. Qed.
Lemma InterpUniv_Join' i j (A B : Tm 0) PA PB : Lemma InterpUniv_Join' n i j (A B : Tm n) PA PB :
A i PA -> A i PA ->
B j PB -> B j PB ->
join A B -> Join.R A B ->
PA = PB. PA = PB.
Proof. Proof.
have [? ?] : i <= max i j /\ j <= max i j by lia. have [? ?] : i <= max i j /\ j <= max i j by lia.
@ -335,16 +471,16 @@ Proof.
eauto using InterpUniv_Join. eauto using InterpUniv_Join.
Qed. Qed.
Lemma InterpUniv_Functional' i j A PA PB : Lemma InterpUniv_Functional' n i j A PA PB :
A i PA -> A : Tm n i PA ->
A j PB -> A j PB ->
PA = PB. PA = PB.
Proof. Proof.
hauto l:on use:InterpUniv_Join', join_refl. hauto l:on use:InterpUniv_Join', join_refl.
Qed. Qed.
Lemma InterpExt_Bind_inv_nopf i I p A B P (h : TBind p A B i ;; I P) : Lemma InterpExt_Bind_inv_nopf i n I p A B P (h : TBind p A B i ;; I P) :
exists (PA : Tm 0 -> Prop), exists (PA : Tm n -> Prop),
A i ;; I PA /\ A i ;; I PA /\
(forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i ;; I PB) /\ (forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i ;; I PB) /\
P = BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i ;; I PB). P = BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i ;; I PB).
@ -365,34 +501,42 @@ Proof.
split; hauto q:on use:InterpExt_Functional. split; hauto q:on use:InterpExt_Functional.
Qed. Qed.
Lemma InterpUniv_Bind_inv_nopf i p A B P (h : TBind p A B i P) : Lemma InterpUniv_Bind_inv_nopf n i p A B P (h : TBind p A B i P) :
exists (PA : Tm 0 -> Prop), exists (PA : Tm n -> Prop),
A i PA /\ A i PA /\
(forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i PB) /\ (forall a, PA a -> exists PB, subst_Tm (scons a VarTm) B i PB) /\
P = BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i PB). P = BindSpace p PA (fun a PB => subst_Tm (scons a VarTm) B i PB).
Proof. hauto l:on use:InterpExt_Bind_inv_nopf rew:db:InterpUniv. Qed. Proof. hauto l:on use:InterpExt_Bind_inv_nopf rew:db:InterpUniv. Qed.
Lemma InterpExt_back_clos i I (A : Tm 0) PA : Lemma InterpExt_back_clos n i I (A : Tm n) PA :
(forall j, forall a b, (RPar.R a b) -> I j b -> I j a) -> (forall j, j < i -> forall a b, (RPar'.R a b) -> I j b -> I j a) ->
A i ;; I PA -> A i ;; I PA ->
forall a b, (RPar.R a b) -> forall a b, (RPar'.R a b) ->
PA b -> PA a. PA b -> PA a.
Proof. Proof.
move => hI h. move => hI h.
elim : A PA /h. elim : A PA /h.
- hauto q:on ctrs:rtc unfold:wne.
- move => p A B PA PF hPA ihPA hTot hRes ihPF a b hr. - move => p A B PA PF hPA ihPA hTot hRes ihPF a b hr.
case : p => //=. case : p => //=.
+ have : forall b0 b1 a, RPar.R b0 b1 -> RPar.R (App b0 a) (App b1 a) + have : forall b0 b1 a, RPar'.R b0 b1 -> RPar'.R (App b0 a) (App b1 a)
by hauto lq:on ctrs:RPar.R use:RPar.refl. by hauto lq:on ctrs:RPar'.R use:RPar'.refl.
hauto lq:on rew:off unfold:ProdSpace. hauto lq:on rew:off unfold:ProdSpace.
+ hauto lq:on ctrs:rtc unfold:SumSpace. + hauto lq:on ctrs:rtc unfold:SumSpace.
- eauto. - eauto.
- eauto. - eauto.
Qed. Qed.
Lemma InterpUniv_back_clos i (A : Tm 0) PA : Lemma InterpExt_back_clos_star n i I (A : Tm n) PA :
(forall j, j < i -> forall a b, (RPar'.R a b) -> I j b -> I j a) ->
A i ;; I PA ->
forall a b, (rtc RPar'.R a b) ->
PA b -> PA a.
Proof. induction 3; hauto l:on use:InterpExt_back_clos. Qed.
Lemma InterpUniv_back_clos n i (A : Tm n) PA :
A i PA -> A i PA ->
forall a b, (RPar.R a b) -> forall a b, (RPar'.R a b) ->
PA b -> PA a. PA b -> PA a.
Proof. Proof.
simp InterpUniv. simp InterpUniv.
@ -400,9 +544,9 @@ Proof.
hauto lq:on ctrs:rtc use:InterpUnivN_back_preservation_star. hauto lq:on ctrs:rtc use:InterpUnivN_back_preservation_star.
Qed. Qed.
Lemma InterpUniv_back_clos_star i (A : Tm 0) PA : Lemma InterpUniv_back_clos_star n i (A : Tm n) PA :
A i PA -> A i PA ->
forall a b, rtc RPar.R a b -> forall a b, rtc RPar'.R a b ->
PA b -> PA a. PA b -> PA a.
Proof. Proof.
move => h a b. move => h a b.
@ -410,30 +554,101 @@ Proof.
hauto lq:on use:InterpUniv_back_clos. hauto lq:on use:InterpUniv_back_clos.
Qed. Qed.
Definition ρ_ok {n} Γ (ρ : fin n -> Tm 0) := forall i m PA, Lemma pars'_wn {n} a b :
subst_Tm ρ (Γ i) m PA -> PA (ρ i). rtc RPar'.R a b ->
@wn n b ->
wn a.
Proof. sfirstorder unfold:wn use:@relations.rtc_transitive. Qed.
Definition SemWt {n} Γ (a A : Tm n) := forall ρ, ρ_ok Γ ρ -> exists m PA, subst_Tm ρ A m PA /\ PA (subst_Tm ρ a). (* P identifies a set of "reducibility candidates" *)
Definition CR {n} (P : Tm n -> Prop) :=
(forall a, P a -> wn a) /\
(forall a, ne a -> P a).
Lemma adequacy_ext i n I A PA
(hI0 : forall j, j < i -> forall a b, (RPar'.R a b) -> I j b -> I j a)
(hI : forall j, j < i -> CR (I j))
(h : A : Tm n i ;; I PA) :
CR PA /\ wn A.
Proof.
elim : A PA / h.
- hauto unfold:wne use:wne_wn.
- move => p A B PA PF hA hPA hTot hRes ihPF.
rewrite /CR.
have hb : PA Bot by firstorder.
repeat split.
+ case : p => /=.
* qauto l:on use:ext_wn unfold:ProdSpace, CR.
* rewrite /SumSpace => a []; first by eauto with nfne.
move => [q0][q1]*.
have : wn q0 /\ wn q1 by hauto q:on.
qauto l:on use:wn_pair, pars'_wn.
+ case : p => /=.
* rewrite /ProdSpace.
move => a ha c PB hc hPB.
have hc' : wn c by sfirstorder.
have : wne (App a c) by hauto lq:on use:wne_app ctrs:rtc.
have h : (forall a, ne a -> PB a) by sfirstorder.
suff : (forall a, wne a -> PB a) by hauto l:on.
move => a0 [a1 [h0 h1]].
eapply InterpExt_back_clos_star with (b := a1); eauto.
* rewrite /SumSpace.
move => a ha. left.
sfirstorder ctrs:rtc.
+ have wnA : wn A by firstorder.
apply wn_bind => //.
apply wn_antirenaming with (ρ := scons Bot VarTm);first by hauto q:on inv:option.
hauto lq:on.
- hauto l:on.
- hauto lq:on rew:off ctrs:rtc.
Qed.
Lemma adequacy i n A PA
(h : A : Tm n i PA) :
CR PA /\ wn A.
Proof.
move : i A PA h.
elim /Wf_nat.lt_wf_ind => i ih A PA.
simp InterpUniv.
apply adequacy_ext.
hauto lq:on ctrs:rtc use:InterpUnivN_back_preservation_star.
hauto l:on use:InterpExt_Ne rew:db:InterpUniv.
Qed.
Lemma adequacy_wne i n A PA a : A : Tm n i PA -> wne a -> PA a.
Proof. qauto l:on use:InterpUniv_back_clos_star, adequacy unfold:CR. Qed.
Lemma adequacy_wn i n A PA (h : A : Tm n i PA) a : PA a -> wn a.
Proof. hauto q:on use:adequacy. Qed.
Definition ρ_ok {n} (Γ : fin n -> Tm n) (ρ : fin n -> Tm 0) := forall i k PA,
subst_Tm ρ (Γ i) k PA -> PA (ρ i).
Definition SemWt {n} Γ (a A : Tm n) := forall ρ, ρ_ok Γ ρ -> exists k PA, subst_Tm ρ A k PA /\ PA (subst_Tm ρ a).
Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70). Notation "Γ ⊨ a ∈ A" := (SemWt Γ a A) (at level 70).
(* Semantic context wellformedness *) (* Semantic context wellformedness *)
Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ Γ i Univ j. Definition SemWff {n} Γ := forall (i : fin n), exists j, Γ Γ i Univ j.
Notation "⊨ Γ" := (SemWff Γ) (at level 70). Notation "⊨ Γ" := (SemWff Γ) (at level 70).
Lemma ρ_ok_nil ρ : Lemma ρ_ok_bot n (Γ : fin n -> Tm n) :
ρ_ok null ρ. ρ_ok Γ (fun _ => Bot).
Proof. rewrite /ρ_ok. inversion i; subst. Qed. Proof.
rewrite /ρ_ok.
hauto q:on use:adequacy.
Qed.
Lemma ρ_ok_cons n i (Γ : fin n -> Tm n) ρ a PA A : Lemma ρ_ok_cons n i (Γ : fin n -> Tm n) ρ a PA A :
subst_Tm ρ A i PA -> PA a -> subst_Tm ρ A i PA -> PA a ->
ρ_ok Γ ρ -> ρ_ok Γ ρ ->
ρ_ok (funcomp (ren_Tm shift) (scons A Γ)) ((scons a ρ)). ρ_ok (funcomp (ren_Tm shift) (scons A Γ)) (scons a ρ).
Proof. Proof.
move => h0 h1 h2. move => h0 h1 h2.
rewrite /ρ_ok. rewrite /ρ_ok.
move => j. move => j.
destruct j as [j|]. destruct j as [j|].
- move => m PA0. asimpl => ?. - move => m PA0. asimpl => ?.
asimpl.
firstorder. firstorder.
- move => m PA0. asimpl => h3. - move => m PA0. asimpl => h3.
have ? : PA0 = PA by eauto using InterpUniv_Functional'. have ? : PA0 = PA by eauto using InterpUniv_Functional'.
@ -455,7 +670,7 @@ Proof.
rewrite /ρ_ok in hρ. rewrite /ρ_ok in hρ.
move => h. move => h.
rewrite /funcomp. rewrite /funcomp.
apply hρ with (m := m'). apply hρ with (k := m').
move : h. rewrite -. move : h. rewrite -.
by asimpl. by asimpl.
Qed. Qed.
@ -480,6 +695,17 @@ Proof.
hauto lq:on inv:option unfold:renaming_ok. hauto lq:on inv:option unfold:renaming_ok.
Qed. Qed.
Lemma SemWt_Wn n Γ (a : Tm n) A :
Γ a A ->
wn a /\ wn A.
Proof.
move => h.
have {}/h := ρ_ok_bot _ Γ => h.
have h0 : wn (subst_Tm (fun _ : fin n => (Bot : Tm 0)) A) by hauto l:on use:adequacy.
have h1 : wn (subst_Tm (fun _ : fin n => (Bot : Tm 0)) a)by hauto l:on use:adequacy_wn.
move {h}. hauto lq:on use:wn_antirenaming.
Qed.
Lemma SemWt_Univ n Γ (A : Tm n) i : Lemma SemWt_Univ n Γ (A : Tm n) i :
Γ A Univ i <-> Γ A Univ i <->
forall ρ, ρ_ok Γ ρ -> exists S, subst_Tm ρ A i S. forall ρ, ρ_ok Γ ρ -> exists S, subst_Tm ρ A i S.
@ -542,12 +768,12 @@ Qed.
Lemma ST_Conv n Γ (a : Tm n) A B i : Lemma ST_Conv n Γ (a : Tm n) A B i :
Γ a A -> Γ a A ->
Γ B Univ i -> Γ B Univ i ->
join A B -> Join.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 : join (subst_Tm ρ A) (subst_Tm ρ B) by eauto using join_substing. have {}h0 : Join.R (subst_Tm ρ A) (subst_Tm ρ B) by eauto using Join.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 ? : PA = S by eauto using InterpUniv_Join'. subst.
@ -572,7 +798,7 @@ Proof.
intros (m & PB0 & hPB0 & hPB0'). intros (m & PB0 & hPB0 & hPB0').
replace PB0 with PB in * by hauto l:on use:InterpUniv_Functional'. replace PB0 with PB in * by hauto l:on use:InterpUniv_Functional'.
apply : InterpUniv_back_clos; eauto. apply : InterpUniv_back_clos; eauto.
apply : RPar.AppAbs'; eauto using RPar.refl. apply : RPar'.AppAbs'; eauto using RPar'.refl.
by asimpl. by asimpl.
Qed. Qed.
@ -604,7 +830,7 @@ Proof.
simpl in hPPi. simpl in hPPi.
move /InterpUniv_Bind_inv_nopf : hPPi. move /InterpUniv_Bind_inv_nopf : hPPi.
move => [PA [hPA [hTot ?]]]. subst=>/=. move => [PA [hPA [hTot ?]]]. subst=>/=.
rewrite /SumSpace. rewrite /SumSpace. right.
exists (subst_Tm ρ a), (subst_Tm ρ b). exists (subst_Tm ρ a), (subst_Tm ρ b).
split. split.
- hauto l:on use:Pars.substing. - hauto l:on use:Pars.substing.
@ -626,24 +852,25 @@ Proof.
move : h0 => [S][h2][h3]?. subst. move : h0 => [S][h2][h3]?. subst.
move : h1 => /=. move : h1 => /=.
rewrite /SumSpace. rewrite /SumSpace.
case; first by hauto lq:on use:adequacy_wne, wne_proj.
move => [a0 [b0 [h4 [h5 h6]]]]. move => [a0 [b0 [h4 [h5 h6]]]].
exists m, S. split => //=. exists m, S. split => //=.
have {}h4 : rtc RPar.R (Proj PL (subst_Tm ρ a)) (Proj PL (Pair a0 b0)) by eauto using RPars.ProjCong. have {}h4 : rtc RPar'.R (Proj PL (subst_Tm ρ a)) (Proj PL (Pair a0 b0)) by eauto using RPars'.ProjCong.
have ? : RPar.R (Proj PL (Pair a0 b0)) a0 by hauto l:on use:RPar.refl, RPar.ProjPair'. have ? : RPar'.R (Proj PL (Pair a0 b0)) a0 by hauto l:on use:RPar'.refl, RPar'.ProjPair'.
have : rtc RPar.R (Proj PL (subst_Tm ρ a)) a0 by eauto using @relations.rtc_r. have : rtc RPar'.R (Proj PL (subst_Tm ρ a)) a0 by eauto using @relations.rtc_r.
move => h. move => h.
apply : InterpUniv_back_clos_star; eauto. apply : InterpUniv_back_clos_star; eauto.
Qed. Qed.
Lemma substing_RPar n m (A : Tm (S n)) ρ (B : Tm m) C : Lemma substing_RPar' n m (A : Tm (S n)) ρ (B : Tm m) C :
RPar.R B C -> RPar'.R B C ->
RPar.R (subst_Tm (scons B ρ) A) (subst_Tm (scons C ρ) A). RPar'.R (subst_Tm (scons B ρ) A) (subst_Tm (scons C ρ) A).
Proof. hauto lq:on inv:option use:RPar.morphing, RPar.refl. Qed. Proof. hauto lq:on inv:option use:RPar'.morphing, RPar'.refl. Qed.
Lemma substing_RPars n m (A : Tm (S n)) ρ (B : Tm m) C : Lemma substing_RPar's n m (A : Tm (S n)) ρ (B : Tm m) C :
rtc RPar.R B C -> rtc RPar'.R B C ->
rtc RPar.R (subst_Tm (scons B ρ) A) (subst_Tm (scons C ρ) A). rtc RPar'.R (subst_Tm (scons B ρ) A) (subst_Tm (scons C ρ) A).
Proof. induction 1; hauto lq:on ctrs:rtc use:substing_RPar. Qed. Proof. induction 1; hauto lq:on ctrs:rtc use:substing_RPar'. Qed.
Lemma ST_Proj2 n Γ (a : Tm n) A B : Lemma ST_Proj2 n Γ (a : Tm n) A B :
Γ a TBind TSig A B -> Γ a TBind TSig A B ->
@ -654,17 +881,156 @@ Proof.
move : h0 => [S][h2][h3]?. subst. move : h0 => [S][h2][h3]?. subst.
move : h1 => /=. move : h1 => /=.
rewrite /SumSpace. rewrite /SumSpace.
move => [a0 [b0 [h4 [h5 h6]]]]. case.
- move => h.
have hp : forall p, wne (Proj p (subst_Tm ρ a)) by auto using wne_proj.
have hp0 := hp PL. have hp1 := hp PR => {hp}.
have : S (Proj PL (subst_Tm ρ a)) by hauto q:on use:adequacy_wne.
move /h3 => [PB]. asimpl. hauto lq:on use:adequacy_wne.
- move => [a0 [b0 [h4 [h5 h6]]]].
specialize h3 with (1 := h5). specialize h3 with (1 := h5).
move : h3 => [PB hPB]. move : h3 => [PB hPB].
have hr : forall p, rtc RPar.R (Proj p (subst_Tm ρ a)) (Proj p (Pair a0 b0)) by eauto using RPars.ProjCong. have hr : forall p, rtc RPar'.R (Proj p (subst_Tm ρ a)) (Proj p (Pair a0 b0)) by eauto using RPars'.ProjCong.
have hrl : RPar.R (Proj PL (Pair a0 b0)) a0 by hauto l:on use:RPar.ProjPair', RPar.refl. have hrl : RPar'.R (Proj PL (Pair a0 b0)) a0 by hauto l:on use:RPar'.ProjPair', RPar'.refl.
have hrr : RPar.R (Proj PR (Pair a0 b0)) b0 by hauto l:on use:RPar.ProjPair', RPar.refl. have hrr : RPar'.R (Proj PR (Pair a0 b0)) b0 by hauto l:on use:RPar'.ProjPair', RPar'.refl.
exists m, PB. exists m, PB.
asimpl. split. asimpl. split.
- have h : rtc RPar.R (Proj PL (subst_Tm ρ a)) a0 by eauto using @relations.rtc_r. + have h : rtc RPar'.R (Proj PL (subst_Tm ρ a)) a0 by eauto using @relations.rtc_r.
have {}h : rtc RPar.R (subst_Tm (scons (Proj PL (subst_Tm ρ a)) ρ) B) (subst_Tm (scons a0 ρ) B) by eauto using substing_RPars. have {}h : rtc RPar'.R (subst_Tm (scons (Proj PL (subst_Tm ρ a)) ρ) B) (subst_Tm (scons a0 ρ) B) by eauto using substing_RPar's.
move : hPB. asimpl. move : hPB. asimpl.
eauto using InterpUnivN_back_preservation_star. eauto using InterpUnivN_back_preservation_star.
- hauto lq:on use:@relations.rtc_r, InterpUniv_back_clos_star. + hauto lq:on use:@relations.rtc_r, InterpUniv_back_clos_star.
Qed. Qed.
Lemma ne_nf_preservation n (a b : Tm n) : ERed.R b a -> (ne a -> ne b) /\ (nf a -> nf b).
Proof.
move => h. elim : n b a /h => //=.
- move => n a.
split => //=.
hauto lqb:on use:ne_nf_ren db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
- hauto lqb:on db:nfne.
Qed.
Fixpoint size_tm {n} (a : Tm n) :=
match a with
| VarTm _ => 1
| TBind _ A B => 1 + Nat.add (size_tm A) (size_tm B)
| Abs a => 1 + size_tm a
| App a b => 1 + Nat.add (size_tm a) (size_tm b)
| Proj p a => 1 + size_tm a
| Pair a b => 1 + Nat.add (size_tm a) (size_tm b)
| Bot => 1
| Const _ => 1
| Univ _ => 1
end.
Lemma size_tm_ren n m (ξ : fin n -> fin m) a : size_tm (ren_Tm ξ a) = size_tm a.
Proof.
move : m ξ. elim : n / a => //=; scongruence.
Qed.
#[export]Hint Rewrite size_tm_ren : size_tm.
Lemma size_η_lt n (a b : Tm n) :
ERed.R b a ->
size_tm b < size_tm a.
Proof.
move => h. elim : b a / h => //=; hauto l:on rew:db:size_tm.
Qed.
Lemma ered_local_confluence n (a b c : Tm n) :
ERed.R b a ->
ERed.R c a ->
exists d, rtc ERed.R d b /\ rtc ERed.R d c.
Proof.
move => h. move : c.
elim : n b a / h => n.
- move => a c.
elim /ERed.inv => //= _.
+ move => ? ? [*]. subst.
have : subst_Tm (scons Bot VarTm) (ren_Tm shift c) = (subst_Tm (scons Bot VarTm) (ren_Tm shift a))
by congruence.
asimpl => ?. subst.
eauto using rtc_refl.
+ move => a0 a1 ha ? [*]. subst.
elim /ERed.inv : ha => //= _.
* move => a1 a2 b0 ha ? [*]. subst.
have [a2 [h0 h1]] : exists a2, ERed.R a2 a /\ a1 = ren_Tm shift a2 by admit. subst.
eexists. split; cycle 1.
apply : relations.rtc_r; cycle 1.
apply ERed.AppEta.
apply rtc_refl.
eauto using relations.rtc_once.
* hauto q:on ctrs:rtc, ERed.R inv:ERed.R.
- move => a c ha.
elim /ERed.inv : ha => //= _.
+ hauto l:on.
+ move => a0 a1 b0 ha ? [*]. subst.
elim /ERed.inv : ha => //= _.
move => p a1 a2 ha ? [*]. subst.
exists a1. split. by apply relations.rtc_once.
apply : rtc_l. apply ERed.PairEta.
apply : rtc_l. apply ERed.PairCong1. eauto using ERed.ProjCong.
apply rtc_refl.
+ move => a0 b0 b1 ha ? [*]. subst.
elim /ERed.inv : ha => //= _ p a0 a1 h ? [*]. subst.
exists a0. split; first by apply relations.rtc_once.
apply : rtc_l; first by apply ERed.PairEta.
apply relations.rtc_once.
hauto lq:on ctrs:ERed.R.
- move => a0 a1 ha iha c.
elim /ERed.inv => //= _.
+ move => a2 ? [*]. subst.
elim /ERed.inv : ha => //=_.
* move => a1 a2 b0 ha ? [*] {iha}. subst.
have [a0 [h0 h1]] : exists a0, ERed.R a0 c /\ a1 = ren_Tm shift a0 by admit. subst.
exists a0. split; last by apply relations.rtc_once.
apply relations.rtc_once. apply ERed.AppEta.
* hauto q:on inv:ERed.R.
+ hauto l:on use:EReds.AbsCong.
- move => a0 a1 b ha iha c.
elim /ERed.inv => //= _.
+ hauto lq:on ctrs:rtc use:EReds.AppCong.
+ hauto lq:on use:@relations.rtc_once ctrs:ERed.R.
- move => a b0 b1 hb ihb c.
elim /ERed.inv => //=_.
+ move => a0 a1 a2 ha ? [*]. subst.
move {ihb}. exists (App a0 b0).
hauto lq:on use:@relations.rtc_once ctrs:ERed.R.
+ hauto lq:on ctrs:rtc use:EReds.AppCong.
- move => a0 a1 b ha iha c.
elim /ERed.inv => //= _.
+ move => ? ?[*]. subst.
elim /ERed.inv : ha => //= _ p a1 a2 h ? [*]. subst.
exists a1. split; last by apply relations.rtc_once.
apply : rtc_l. apply ERed.PairEta.
apply relations.rtc_once. hauto lq:on ctrs:ERed.R.
+ hauto lq:on ctrs:rtc use:EReds.PairCong.
+ hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
- move => a b0 b1 hb hc c. elim /ERed.inv => //= _.
+ move => ? ? [*]. subst.
elim /ERed.inv : hb => //= _ p a0 a1 ha ? [*]. subst.
move {hc}.
exists a0. split; last by apply relations.rtc_once.
apply : rtc_l; first by apply ERed.PairEta.
hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+ hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+ hauto lq:on ctrs:rtc use:EReds.PairCong.
- qauto l:on inv:ERed.R use:EReds.ProjCong.
- move => p A0 A1 B hA ihA.
move => c. elim/ERed.inv => //=.
+ hauto lq:on ctrs:rtc use:EReds.BindCong.
+ hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
- move => p A B0 B1 hB ihB c.
elim /ERed.inv => //=.
+ hauto lq:on ctrs:ERed.R use:@relations.rtc_once.
+ hauto lq:on ctrs:rtc use:EReds.BindCong.
Admitted.