Add compile

This commit is contained in:
Yiyun Liu 2025-01-20 15:33:46 -05:00
parent f3718707f2
commit d68df5d0bc
2 changed files with 172 additions and 13 deletions

View file

@ -410,16 +410,16 @@ Module RPar.
qauto l:on ctrs:R inv:option.
Qed.
Lemma var_or_bot_imp {n} (a b : Tm n) :
var_or_bot a ->
a = b -> ~~ var_or_bot b -> False.
Lemma var_or_const_imp {n} (a b : Tm n) :
var_or_const a ->
a = b -> ~~ var_or_const b -> False.
Proof.
hauto lq:on inv:Tm.
Qed.
Lemma var_or_bot_up n m (ρ : fin n -> Tm m) :
(forall i, var_or_bot (ρ i)) ->
(forall i, var_or_bot (up_Tm_Tm ρ i)).
Lemma var_or_const_up n m (ρ : fin n -> Tm m) :
(forall i, var_or_const (ρ i)) ->
(forall i, var_or_const (up_Tm_Tm ρ i)).
Proof.
move => h /= [i|].
- asimpl.
@ -430,10 +430,10 @@ Module RPar.
- sfirstorder.
Qed.
Local Ltac antiimp := qauto l:on use:var_or_bot_imp.
Local Ltac antiimp := qauto l:on use:var_or_const_imp.
Lemma antirenaming n m (a : Tm n) (b : Tm m) (ρ : fin n -> Tm m) :
(forall i, var_or_bot (ρ i)) ->
(forall i, var_or_const (ρ i)) ->
R (subst_Tm ρ a) b -> exists b0, R a b0 /\ subst_Tm ρ b0 = b.
Proof.
move E : (subst_Tm ρ a) => u hρ h.
@ -444,7 +444,7 @@ Module RPar.
case : c => //=; first by antiimp.
move => c [?]. subst.
spec_refl.
have /var_or_bot_up hρ' := hρ.
have /var_or_const_up hρ' := hρ.
move : iha hρ' => /[apply] iha.
move : ihb hρ => /[apply] ihb.
spec_refl.
@ -470,7 +470,7 @@ Module RPar.
- move => n p a0 a1 ha iha m ρ hρ []//=;
first by antiimp.
move => p0 []//= t [*]; first by antiimp. subst.
have /var_or_bot_up {}/iha := hρ => iha.
have /var_or_const_up {}/iha := hρ => iha.
spec_refl. move : iha => [b0 [? ?]]. subst.
eexists. split. apply ProjAbs; eauto. by asimpl.
- move => n p a0 a1 b0 b1 ha iha hb ihb m ρ hρ []//=;
@ -488,7 +488,7 @@ Module RPar.
hauto l:on.
- move => n a0 a1 ha iha m ρ hρ []//=; first by antiimp.
move => t [*]. subst.
have /var_or_bot_up {}/iha := hρ => iha.
have /var_or_const_up {}/iha := hρ => iha.
spec_refl.
move :iha => [b0 [? ?]]. subst.
eexists. split. by apply AbsCong; eauto.
@ -524,7 +524,7 @@ Module RPar.
first by antiimp.
move => ? t t0 [*]. subst.
have {}/iha := (hρ) => iha.
have /var_or_bot_up {}/ihB := (hρ) => ihB.
have /var_or_const_up {}/ihB := (hρ) => ihB.
spec_refl.
move : iha => [b0 [? ?]].
move : ihB => [c0 [? ?]]. subst.
@ -1838,7 +1838,7 @@ Definition prov_extract_spec {n} u (a : Tm n) :=
| TBind p A B => exists A0 B0, extract a = TBind p A0 B0 /\ rtc Par.R A A0 /\ rtc Par.R B B0
| Univ i => extract a = Univ i
| VarTm i => extract a = VarTm i
| (Const TPi) => extract a = (Const TPi)
| (Const i) => extract a = (Const i)
| _ => True
end.
@ -2250,6 +2250,15 @@ Proof.
apply prov_extract.
Qed.
Lemma pars_const_inv n i (c : Tm n) :
rtc Par.R (Const i) c ->
extract c = Const i.
Proof.
have : prov (Const i) (Const i : Tm n) by sfirstorder.
move : prov_pars. repeat move/[apply].
apply prov_extract.
Qed.
Lemma pars_pi_inv n p (A : Tm n) B C :
rtc Par.R (TBind p A B) C ->
exists A0 B0, extract C = TBind p A0 B0 /\
@ -2277,6 +2286,14 @@ Proof.
sauto l:on use:pars_univ_inv.
Qed.
Lemma pars_const_inj n i j (C : Tm n) :
rtc Par.R (Const i) C ->
rtc Par.R (Const j) C ->
i = j.
Proof.
sauto l:on use:pars_const_inv.
Qed.
Lemma pars_pi_inj n p0 p1 (A0 A1 : Tm n) B0 B1 C :
rtc Par.R (TBind p0 A0 B0) C ->
rtc Par.R (TBind p1 A1 B1) C ->
@ -2314,6 +2331,12 @@ Proof.
sfirstorder use:pars_univ_inj.
Qed.
Lemma join_const_inj n i j :
join (Const i : Tm n) (Const j) -> i = j.
Proof.
sfirstorder use:pars_const_inj.
Qed.
Lemma join_pi_inj n p0 p1 (A0 A1 : Tm n) B0 B1 :
join (TBind p0 A0 B0) (TBind p1 A1 B1) ->
p0 = p1 /\ join A0 A1 /\ join B0 B1.