Add injectivity lemma for abs
This commit is contained in:
parent
300530a93f
commit
9bd554b513
2 changed files with 34 additions and 0 deletions
|
@ -681,6 +681,7 @@ Proof.
|
||||||
suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
|
suff [v0 [hv00 hv01]] : exists v0, rtc ERed.R va' v0 /\ rtc ERed.R vb' v0.
|
||||||
repeat split =>//=. sfirstorder.
|
repeat split =>//=. sfirstorder.
|
||||||
simpl in *; by lia.
|
simpl in *; by lia.
|
||||||
|
|
||||||
admit.
|
admit.
|
||||||
+ case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
|
+ case : b => //=; try qauto depth:1 use:T_AbsPair_Imp, T_PairBind_Imp, T_PairUniv_Imp.
|
||||||
move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
|
move => a1 b1 a0 b0 h _ _ Γ A hu0 hu1.
|
||||||
|
|
|
@ -2319,6 +2319,14 @@ Module DJoin.
|
||||||
hauto lq:on rew:off ctrs:rtc unfold:R use:REReds.substing.
|
hauto lq:on rew:off ctrs:rtc unfold:R use:REReds.substing.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Lemma renaming n m (a b : PTm n) (ρ : fin n -> fin m) :
|
||||||
|
R a b -> R (ren_PTm ρ a) (ren_PTm ρ b).
|
||||||
|
Proof. substify. apply substing. Qed.
|
||||||
|
|
||||||
|
Lemma weakening n (a b : PTm n) :
|
||||||
|
R a b -> R (ren_PTm shift a) (ren_PTm shift b).
|
||||||
|
Proof. apply renaming. Qed.
|
||||||
|
|
||||||
Lemma cong n m (a : PTm (S n)) c d (ρ : fin n -> PTm m) :
|
Lemma cong n m (a : PTm (S n)) c d (ρ : fin n -> PTm m) :
|
||||||
R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) a).
|
R c d -> R (subst_PTm (scons c ρ) a) (subst_PTm (scons d ρ) a).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -2327,6 +2335,31 @@ Module DJoin.
|
||||||
hauto q:on ctrs:rtc inv:option use:REReds.cong.
|
hauto q:on ctrs:rtc inv:option use:REReds.cong.
|
||||||
Qed.
|
Qed.
|
||||||
|
|
||||||
|
Lemma abs_inj n (a0 a1 : PTm (S n)) :
|
||||||
|
SN a0 -> SN a1 ->
|
||||||
|
R (PAbs a0) (PAbs a1) ->
|
||||||
|
R a0 a1.
|
||||||
|
Proof.
|
||||||
|
move => sn0 sn1.
|
||||||
|
move /weakening => /=.
|
||||||
|
move /AppCong. move /(_ (VarPTm var_zero) (VarPTm var_zero)).
|
||||||
|
move => /(_ ltac:(sfirstorder use:refl)).
|
||||||
|
move => h.
|
||||||
|
have /FromRRed1 h0 : RRed.R (PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a0)) (VarPTm var_zero)) a0 by apply RRed.AppAbs'; asimpl.
|
||||||
|
have /FromRRed0 h1 : RRed.R (PApp (PAbs (ren_PTm (upRen_PTm_PTm shift) a1)) (VarPTm var_zero)) a1 by apply RRed.AppAbs'; asimpl.
|
||||||
|
have sn0' := sn0.
|
||||||
|
have sn1' := sn1.
|
||||||
|
eapply sn_renaming with (ξ := (upRen_PTm_PTm shift)) in sn0.
|
||||||
|
eapply sn_renaming with (ξ := (upRen_PTm_PTm shift)) in sn1.
|
||||||
|
apply : transitive; try apply h0.
|
||||||
|
apply : N_Exp. apply N_β. sauto.
|
||||||
|
by asimpl.
|
||||||
|
apply : transitive; try apply h1.
|
||||||
|
apply : N_Exp. apply N_β. sauto.
|
||||||
|
by asimpl.
|
||||||
|
eauto.
|
||||||
|
Qed.
|
||||||
|
|
||||||
End DJoin.
|
End DJoin.
|
||||||
|
|
||||||
|
|
||||||
|
|
Loading…
Add table
Add a link
Reference in a new issue