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=24dcd4c1f160514bec29ac8e273ce33aaebef480;hb=a09b60bd574adf1a7d3e423023009cb20c79d449;hp=3d547a49cb67dbde8ad4736fd1b5014db1a02909;hpb=f1b5681cc8b74b016c7287d901564547fa70cdbd;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 3d547a49c..24dcd4c1f 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 @@ -43,7 +43,7 @@ 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 +#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-. @@ -52,13 +52,28 @@ 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 +#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 ************************************************) (* Basic_2A1: uses: lsx_inv_flat *) 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-.