X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcpxs_etc.ma;h=0340000b950e31630a5a420bbb84a9941fb6b43e;hb=c3832abc23bb0907df2deb6751f4a46d213675b7;hp=a28b01e54be3195e3a4b3c2a9b39977f38a959c9;hpb=228776cd21ca563e6dc44bfdf7746e69dc7b66d6;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma index a28b01e54..0340000b9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_etc.ma @@ -5,21 +5,6 @@ include "basic_2/rt_computation/cpxs_cpxs.ma". (* Properties on sn extended parallel reduction for local environments ******) -lemma lpx_cpx_trans: ∀h,G. b_c_transitive … (cpx h G) (λ_.lpx h G). -#h #G #L2 #T1 #T2 #HT12 elim HT12 -G -L2 -T1 -T2 -[ /2 width=3 by/ -| /3 width=2 by cpx_cpxs, cpx_st/ -| #I #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12 - elim (lpx_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H - elim (lpx_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct - /4 width=7 by cpxs_delta, cpxs_strap2/ -|4,9: /4 width=1 by cpxs_beta, cpxs_bind, lpx_pair/ -|5,7,8: /3 width=1 by cpxs_flat, cpxs_ct, cpxs_eps/ -| /4 width=3 by cpxs_zeta, lpx_pair/ -| /4 width=3 by cpxs_theta, cpxs_strap1, lpx_pair/ -] -qed-. - lemma cpx_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h, o] V2 → ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈[h, o] T2 → ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ⬈*[h, o] ⓑ{a,I}V2.T2. @@ -27,10 +12,6 @@ lemma cpx_bind2: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h, o] V2 → (* Advanced properties ******************************************************) -lemma lpx_cpxs_trans: ∀h,o,G. b_rs_transitive … (cpx h o G) (λ_.lpx h o G). -#h #o #G @b_c_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *) -qed-. - lemma cpxs_bind2_dx: ∀h,o,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h, o] V2 → ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ⬈*[h, o] T2 → ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ⬈*[h, o] ⓑ{a,I}V2.T2.