X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx_cnx_vector.ma;h=d0072aa83cb4e80693abec740f366ed6d29213be;hp=4a93b473e1c6a0f0e9ac330f3fdcc0149632f12d;hb=4173283e148199871d787c53c0301891deb90713;hpb=a67fc50ccfda64377e2c94c18c3a0d9265f651db diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cnx_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cnx_vector.ma index 4a93b473e..d0072aa83 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cnx_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_cnx_vector.ma @@ -23,34 +23,20 @@ include "basic_2/rt_computation/csx_vector.ma". (* Properties with normal terms for unbound parallel rt-transition **********) (* Basic_1: was just: sn3_appls_lref *) -lemma csx_applv_cnx: ∀h,o,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄ → - ∀Vs. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃Vs⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.T⦄. -#h #o #G #L #T #H1T #H2T #Vs elim Vs -Vs -[ #_ normalize in ⊢ (?????%); /2 width=1/ +lemma csx_applv_cnx: ∀h,G,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ⬈[h] 𝐍⦃T⦄ → + ∀Vs. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃Vs⦄ → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⒶVs.T⦄. +#h #G #L #T #H1T #H2T #Vs elim Vs -Vs +[ #_ normalize in ⊢ (????%); /2 width=1 by cnx_csx/ | #V #Vs #IHV #H elim (csxv_inv_cons … H) -H #HV #HVs @csx_appl_simple_theq /2 width=1 by applv_simple/ -IHV -HV -HVs #X #H #H0 - lapply (cpxs_fwd_cnx_vector … o … H) -H // -H1T -H2T #H + lapply (cpxs_fwd_cnx_vector … H) -H // -H1T -H2T #H elim (H0) -H0 // ] qed. (* Advanced properties ******************************************************) -lemma csx_applv_sort: ∀h,o,G,L,s,Vs. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃Vs⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⒶVs.⋆s⦄. -#h #o #G #L #s elim (deg_total h o s) -#d generalize in match s; -s elim d -d -[ /3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ -| #d #IHd #s #Hd #Vs elim Vs -Vs /2 width=1 by/ - #V #Vs #IHVs #HVVs - elim (csxv_inv_cons … HVVs) #HV #HVs - @csx_appl_simple_theq /2 width=1 by applv_simple, simple_atom/ - #X #H #H0 - elim (cpxs_fwd_sort_vector … o … H) -H #H - [ elim H0 -H0 // - | -H0 @(csx_cpxs_trans … (Ⓐ(V⨮Vs).⋆(next h s))) - /3 width=1 by cpxs_flat_dx, deg_next_SO/ - ] -] -qed. +lemma csx_applv_sort: ∀h,G,L,s,Vs. ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃Vs⦄ → ⦃G, L⦄ ⊢ ⬈*[h] 𝐒⦃ⒶVs.⋆s⦄. +/3 width=6 by csx_applv_cnx, cnx_sort, simple_atom/ qed.