X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcrx_append.ma;h=56667b20f28db966fb14943f3b5b87c82f2eafcf;hb=1f30483032488ac4df2310b68fe8146e05524fec;hp=bd91ad1d5c1f2ebd5250605ee9b91e7ed07d6759;hpb=29973426e0227ee48368d1c24dc0c17bf2baef77;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma index bd91ad1d5..56667b20f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma @@ -15,17 +15,18 @@ include "basic_2/relocation/ldrop_append.ma". include "basic_2/reduction/crx.ma". -(* CONTEXT-SENSITIVE EXTENDED REDUCIBLE TERMS *******************************) +(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************) (* Advanved properties ******************************************************) -lemma crx_append_sn: ∀h,g,L,K,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⦃h, K @@ L⦄ ⊢ 𝐑[h, g]⦃T⦄. -#h #g #L #K0 #T #H elim H -L -T /2 width=1/ /2 width=2/ +lemma crx_append_sn: ∀h,g,G,L,K,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ⦃G, K @@ L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄. +#h #g #G #L #K0 #T #H elim H -L -T +/2 width=2 by crx_sort, crx_appl_sn, crx_appl_dx, crx_ri2, crx_ib2_sn, crx_ib2_dx, crx_beta, crx_theta/ #I #L #K #V #i #HLK lapply (ldrop_fwd_length_lt2 … HLK) #Hi -lapply (ldrop_O1_append_sn_le … HLK … K0) -HLK /2 width=2/ -Hi /2 width=4/ +lapply (ldrop_O1_append_sn_le … HLK … K0) -HLK /2 width=4 by crx_delta, lt_to_le/ qed. -lemma trx_crx: ∀h,g,L,T. ⦃h, ⋆⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄. -#h #g #L #T #H lapply (crx_append_sn … H) // +lemma trx_crx: ∀h,g,G,L,T. ⦃G, ⋆⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄. +#h #g #G #L #T #H lapply (crx_append_sn … H) // qed.