X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfsx_lfsx.ma;h=acb4447e3ea90fe82511b2258e9945569d5c3197;hb=47a745462a714af9d65cea7b61af56524bd98fa1;hp=9664ec0a3256d559b960a7c2c2afa4f593d2f26b;hpb=58ea181757dce19b875b2f5a224fe193b2263004;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma index 9664ec0a3..acb4447e3 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/static/lfdeq_lfdeq.ma". include "basic_2/rt_transition/lfpx_lfdeq.ma". include "basic_2/rt_computation/lfsx.ma". @@ -22,7 +21,7 @@ include "basic_2/rt_computation/lfsx.ma". (* Basic_2A1: uses: lsx_lleq_trans *) lemma lfsx_lfdeq_trans: ∀h,o,G,L1,T. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → - ∀L2. L1 ≡[h, o, T] L2 → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄. + ∀L2. L1 ≛[h, o, T] L2 → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄. #h #o #G #L1 #T #H @(lfsx_ind … H) -L1 #L1 #_ #IHL1 #L2 #HL12 @lfsx_intro #L #HL2 #HnL2 elim (lfdeq_lfpx_trans … HL2 … HL12) -HL2 @@ -43,8 +42,8 @@ lemma lfsx_fwd_pair_sn: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ②{I}V.T] 𝐒⦃L G ⊢ ⬈*[h, o, V] 𝐒⦃L⦄. #h #o #I #G #L #V #T #H @(lfsx_ind … H) -L #L1 #_ #IHL1 @lfsx_intro -#L2 #H #HnL12 elim (lfpx_pair_sn_split … o I … T H) -H -/6 width=4 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_pair_sn/ +#L2 #H #HnL12 elim (lfpx_pair_sn_split … H o I T) -H +/6 width=3 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_pair_sn/ qed-. (* Basic_2A1: uses: lsx_fwd_flat_dx *) @@ -52,8 +51,18 @@ lemma lfsx_fwd_flat_dx: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓕ{I}V.T] 𝐒⦃L G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. #h #o #I #G #L #V #T #H @(lfsx_ind … H) -L #L1 #_ #IHL1 @lfsx_intro -#L2 #H #HnL12 elim (lfpx_flat_dx_split … o I … V … H) -H -/6 width=4 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_flat_dx/ +#L2 #H #HnL12 elim (lfpx_flat_dx_split … H o I V) -H +/6 width=3 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_flat_dx/ +qed-. + +(* Basic_2A1: uses: lsx_fwd_bind_dx *) +(* Note: the exclusion binder (ⓧ) makes this more elegant and much simpler *) +lemma lfsx_fwd_bind_dx: ∀h,o,p,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓑ{p,I}V.T] 𝐒⦃L⦄ → + G ⊢ ⬈*[h, o, T] 𝐒⦃L.ⓧ⦄. +#h #o #p #I #G #L #V #T #H @(lfsx_ind … H) -L +#L1 #_ #IH @lfsx_intro +#L2 #H #HT elim (lfpx_bind_dx_split_void … H o p I V) -H +/6 width=5 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_bind_dx_void/ qed-. (* Advanced inversion lemmas ************************************************) @@ -62,3 +71,8 @@ qed-. lemma lfsx_inv_flat: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓕ{I}V.T] 𝐒⦃L⦄ → G ⊢ ⬈*[h, o, V] 𝐒⦃L⦄ ∧ G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. /3 width=3 by lfsx_fwd_pair_sn, lfsx_fwd_flat_dx, conj/ qed-. + +(* Basic_2A1: uses: lsx_inv_bind *) +lemma lfsx_inv_bind: ∀h,o,p,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓑ{p,I}V.T] 𝐒⦃L⦄ → + G ⊢ ⬈*[h, o, V] 𝐒⦃L⦄ ∧ G ⊢ ⬈*[h, o, T] 𝐒⦃L.ⓧ⦄. +/3 width=4 by lfsx_fwd_pair_sn, lfsx_fwd_bind_dx, conj/ qed-.