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=7e1d5cffb36d200d89b1416f2f24a78d0be4d510;hb=670ad7822d59e598a38d9037d482d3de188b170c;hp=7ca826d116941f719be727e7bf8f347191509bde;hpb=7412538ab43afe9a19c5f4be369bed82b2ab6193;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 7ca826d11..7e1d5cffb 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 @@ -20,6 +20,7 @@ include "basic_2/rt_computation/lfsx.ma". (* Advanced properties ******************************************************) +(* Basic_2A1: was just: 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⦄. #h #o #G #L1 #T #H @(lfsx_ind … H) -L1 @@ -28,6 +29,13 @@ lemma lfsx_lfdeq_trans: ∀h,o,G,L1,T. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → /4 width=5 by lfdeq_repl/ qed-. +(* Basic_2A1: was: lsx_lpx_trans *) +lemma lfsx_lfpx_trans: ∀h,o,G,L1,T. G ⊢ ⬈*[h, o, T] 𝐒⦃L1⦄ → + ∀L2. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → G ⊢ ⬈*[h, o, T] 𝐒⦃L2⦄. +#h #o #G #L1 #T #H @(lfsx_ind … H) -L1 #L1 #HL1 #IHL1 #L2 #HL12 +elim (lfdeq_dec h o L1 L2 T) /3 width=4 by lfsx_lfdeq_trans, lfxs_refl/ +qed-. + (* Advanced forward lemmas **************************************************) (* Basic_2A1: includes: lsx_fwd_bind_sn lsx_fwd_flat_sn *) @@ -40,7 +48,6 @@ lemma lfsx_fwd_pair_sn: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ②{I}V.T] 𝐒⦃L /6 width=4 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_pair_sn/ qed-. - (* Basic_2A1: was: lsx_fwd_flat_dx *) lemma lfsx_fwd_flat_dx: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓕ{I}V.T] 𝐒⦃L⦄ → G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄.