Finish renaming

This commit is contained in:
Yiyun Liu 2025-01-24 14:58:35 -07:00
parent 255bd4acbf
commit 398a18d770

View file

@ -2133,26 +2133,16 @@ Lemma pars_univ_inv n i (c : PTm n) :
rtc Par.R (PUniv i) c ->
extract c = PUniv i.
Proof.
have : prov (PUniv i) (Univ i : PTm n) by sfirstorder.
have : prov (PUniv i) (PUniv i : PTm n) by sfirstorder.
move : prov_pars. repeat move/[apply].
apply prov_extract.
Qed.
Lemma pars_const_inv n i (c : PTm n) :
rtc Par.R (Const i) c ->
extract c = Const i.
rtc Par.R (PConst i) c ->
extract c = PConst i.
Proof.
have : prov (Const i) (Const i : PTm n) by sfirstorder.
move : prov_pars. repeat move/[apply].
apply prov_extract.
Qed.
Lemma pars_pi_inv n p (A : PTm n) B C :
rtc Par.R (TBind p A B) C ->
exists A0 B0, extract C = TBind p A0 B0 /\
rtc Par.R A A0 /\ rtc Par.R B B0.
Proof.
have : prov (TBind p A B) (TBind p A B) by hauto lq:on ctrs:prov, rtc.
have : prov (PConst i) (PConst i : PTm n) by sfirstorder.
move : prov_pars. repeat move/[apply].
apply prov_extract.
Qed.
@ -2167,32 +2157,21 @@ Proof.
Qed.
Lemma pars_univ_inj n i j (C : PTm n) :
rtc Par.R (Univ i) C ->
rtc Par.R (Univ j) C ->
rtc Par.R (PUniv i) C ->
rtc Par.R (PUniv j) C ->
i = j.
Proof.
sauto l:on use:pars_univ_inv.
Qed.
Lemma pars_const_inj n i j (C : PTm n) :
rtc Par.R (Const i) C ->
rtc Par.R (Const j) C ->
rtc Par.R (PConst i) C ->
rtc Par.R (PConst j) C ->
i = j.
Proof.
sauto l:on use:pars_const_inv.
Qed.
Lemma pars_pi_inj n p0 p1 (A0 A1 : PTm n) B0 B1 C :
rtc Par.R (TBind p0 A0 B0) C ->
rtc Par.R (TBind p1 A1 B1) C ->
exists A2 B2, p1 = p0 /\ rtc Par.R A0 A2 /\ rtc Par.R A1 A2 /\
rtc Par.R B0 B2 /\ rtc Par.R B1 B2.
Proof.
move /pars_pi_inv => [A2 [B2 [? [h0 h1]]]].
move /pars_pi_inv => [A3 [B3 [? [h2 h3]]]].
exists A2, B2. hauto l:on.
Qed.
Definition join {n} (a b : PTm n) :=
exists c, rtc Par.R a c /\ rtc Par.R b c.
@ -2214,26 +2193,17 @@ Lemma join_refl n (a : PTm n) : join a a.
Proof. hauto lq:on ctrs:rtc unfold:join. Qed.
Lemma join_univ_inj n i j :
join (Univ i : PTm n) (Univ j) -> i = j.
join (PUniv i : PTm n) (PUniv j) -> i = j.
Proof.
sfirstorder use:pars_univ_inj.
Qed.
Lemma join_const_inj n i j :
join (Const i : PTm n) (Const j) -> i = j.
join (PConst i : PTm n) (PConst j) -> i = j.
Proof.
sfirstorder use:pars_const_inj.
Qed.
Lemma join_pi_inj n p0 p1 (A0 A1 : PTm n) B0 B1 :
join (TBind p0 A0 B0) (TBind p1 A1 B1) ->
p0 = p1 /\ join A0 A1 /\ join B0 B1.
Proof.
move => [c []].
move : pars_pi_inj; repeat move/[apply].
sfirstorder unfold:join.
Qed.
Lemma join_substing n m (a b : PTm n) (ρ : fin n -> PTm m) :
join a b ->
join (subst_PTm ρ a) (subst_PTm ρ b).
@ -2242,26 +2212,24 @@ Proof. hauto lq:on unfold:join use:Pars.substing. Qed.
Fixpoint ne {n} (a : PTm n) :=
match a with
| VarPTm i => true
| TBind _ A B => false
| App a b => ne a && nf b
| Abs a => false
| Univ _ => false
| Proj _ a => ne a
| Pair _ _ => false
| Const _ => false
| Bot => true
| PApp a b => ne a && nf b
| PAbs a => false
| PUniv _ => false
| PProj _ a => ne a
| PPair _ _ => false
| PConst _ => false
| PBot => true
end
with nf {n} (a : PTm n) :=
match a with
| VarPTm i => true
| TBind _ A B => nf A && nf B
| App a b => ne a && nf b
| Abs a => nf a
| Univ _ => true
| Proj _ a => ne a
| Pair a b => nf a && nf b
| Const _ => true
| Bot => true
| PApp a b => ne a && nf b
| PAbs a => nf a
| PUniv _ => true
| PProj _ a => ne a
| PPair a b => nf a && nf b
| PConst _ => true
| PBot => true
end.
Lemma ne_nf n a : @ne n a -> nf a.
@ -2290,37 +2258,30 @@ Proof.
Qed.
Lemma wne_app n (a b : PTm n) :
wne a -> wn b -> wne (App a b).
wne a -> wn b -> wne (PApp a b).
Proof.
move => [a0 [? ?]] [b0 [? ?]].
exists (App a0 b0). hauto b:on drew:off use:RPars'.AppCong.
exists (PApp a0 b0). hauto b:on drew:off use:RPars'.AppCong.
Qed.
Lemma wn_abs n a (h : wn a) : @wn n (Abs a).
Lemma wn_abs n a (h : wn a) : @wn n (PAbs a).
Proof.
move : h => [v [? ?]].
exists (Abs v).
exists (PAbs v).
eauto using RPars'.AbsCong.
Qed.
Lemma wn_bind n p A B : wn A -> wn B -> wn (@TBind n p A B).
Proof.
move => [A0 [? ?]] [B0 [? ?]].
exists (TBind p A0 B0).
hauto lqb:on use:RPars'.BindCong.
Qed.
Lemma wn_pair n (a b : PTm n) : wn a -> wn b -> wn (Pair a b).
Lemma wn_pair n (a b : PTm n) : wn a -> wn b -> wn (PPair a b).
Proof.
move => [a0 [? ?]] [b0 [? ?]].
exists (Pair a0 b0).
exists (PPair a0 b0).
hauto lqb:on use:RPars'.PairCong.
Qed.
Lemma wne_proj n p (a : PTm n) : wne a -> wne (Proj p a).
Lemma wne_proj n p (a : PTm n) : wne a -> wne (PProj p a).
Proof.
move => [a0 [? ?]].
exists (Proj p a0). hauto lqb:on use:RPars'.ProjCong.
exists (PProj p a0). hauto lqb:on use:RPars'.ProjCong.
Qed.
Create HintDb nfne.
@ -2348,10 +2309,10 @@ Proof.
Qed.
Lemma ext_wn n (a : PTm n) :
wn (App a Bot) ->
wn (PApp a PBot) ->
wn a.
Proof.
move E : (App a (Bot)) => a0 [v [hr hv]].
move E : (PApp a (PBot)) => a0 [v [hr hv]].
move : a E.
move : hv.
elim : a0 v / hr.
@ -2362,9 +2323,9 @@ Proof.
case : a0 hr0=>// => b0 b1.
elim /RPar'.inv=>// _.
+ move => a0 a3 b2 b3 ? ? [? ?] ? [? ?]. subst.
have ? : b3 = (Bot) by hauto lq:on inv:RPar'.R. subst.
suff : wn (Abs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
have : wn (subst_PTm (scons (Bot) VarPTm) a3) by sfirstorder.
have ? : b3 = (PBot) by hauto lq:on inv:RPar'.R. subst.
suff : wn (PAbs a3) by hauto lq:on ctrs:RPar'.R, rtc unfold:wn.
have : wn (subst_PTm (scons (PBot) VarPTm) a3) by sfirstorder.
move => h. apply wn_abs.
move : h. apply wn_antirenaming.
hauto lq:on rew:off inv:option.
@ -2374,24 +2335,24 @@ Qed.
Module Join.
Lemma ProjCong p n (a0 a1 : PTm n) :
join a0 a1 ->
join (Proj p a0) (Proj p a1).
join (PProj p a0) (PProj p a1).
Proof. hauto lq:on use:Pars.ProjCong unfold:join. Qed.
Lemma PairCong n (a0 a1 b0 b1 : PTm n) :
join a0 a1 ->
join b0 b1 ->
join (Pair a0 b0) (Pair a1 b1).
join (PPair a0 b0) (PPair a1 b1).
Proof. hauto lq:on use:Pars.PairCong unfold:join. Qed.
Lemma AppCong n (a0 a1 b0 b1 : PTm n) :
join a0 a1 ->
join b0 b1 ->
join (App a0 b0) (App a1 b1).
join (PApp a0 b0) (PApp a1 b1).
Proof. hauto lq:on use:Pars.AppCong. Qed.
Lemma AbsCong n (a b : PTm (S n)) :
join a b ->
join (Abs a) (Abs b).
join (PAbs a) (PAbs b).
Proof. hauto lq:on use:Pars.AbsCong. Qed.
Lemma renaming n m (a b : PTm n) (ξ : fin n -> fin m) :
@ -2415,11 +2376,11 @@ Module Join.
End Join.
Lemma abs_eq n a (b : PTm n) :
join (Abs a) b <-> join a (App (ren_PTm shift b) (VarPTm var_zero)).
join (PAbs a) b <-> join a (PApp (ren_PTm shift b) (VarPTm var_zero)).
Proof.
split.
- move => /Join.weakening h.
have {h} : join (App (ren_PTm shift (Abs a)) (VarPTm var_zero)) (App (ren_PTm shift b) (VarPTm var_zero))
have {h} : join (PApp (ren_PTm shift (PAbs a)) (VarPTm var_zero)) (PApp (ren_PTm shift b) (VarPTm var_zero))
by hauto l:on use:Join.AppCong, join_refl.
simpl.
move => ?. apply : join_transitive; eauto.
@ -2431,12 +2392,12 @@ Proof.
Qed.
Lemma pair_eq n (a0 a1 b : PTm n) :
join (Pair a0 a1) b <-> join a0 (Proj PL b) /\ join a1 (Proj PR b).
join (PPair a0 a1) b <-> join a0 (PProj PL b) /\ join a1 (PProj PR b).
Proof.
split.
- move => h.
have /Join.ProjCong {}h := h.
have h0 : forall p, join (if p is PL then a0 else a1) (Proj p (Pair a0 a1))
have h0 : forall p, join (if p is PL then a0 else a1) (PProj p (PPair a0 a1))
by hauto lq:on use:join_symmetric, Join.FromPar, Par.ProjPair', Par.refl.
hauto lq:on rew:off use:join_transitive, join_symmetric.
- move => [h0 h1].
@ -2447,11 +2408,11 @@ Proof.
Qed.
Lemma join_pair_inj n (a0 a1 b0 b1 : PTm n) :
join (Pair a0 a1) (Pair b0 b1) <-> join a0 b0 /\ join a1 b1.
join (PPair a0 a1) (PPair b0 b1) <-> join a0 b0 /\ join a1 b1.
Proof.
split; last by hauto lq:on use:Join.PairCong.
move /pair_eq => [h0 h1].
have : join (Proj PL (Pair b0 b1)) b0 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
have : join (Proj PR (Pair b0 b1)) b1 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
have : join (PProj PL (PPair b0 b1)) b0 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
have : join (PProj PR (PPair b0 b1)) b1 by hauto lq:on use:Join.FromPar, Par.refl, Par.ProjPair'.
eauto using join_transitive.
Qed.