X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fetc%2Fappend%2Fcrx_append.etc;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fetc%2Fappend%2Fcrx_append.etc;h=56667b20f28db966fb14943f3b5b87c82f2eafcf;hb=1fd63df4c77f5c24024769432ea8492748b4ac79;hp=0000000000000000000000000000000000000000;hpb=277fc8ff21ce3dbd6893b1994c55cf5c06a98355;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/etc/append/crx_append.etc b/matita/matita/contribs/lambdadelta/basic_2A/etc/append/crx_append.etc new file mode 100644 index 000000000..56667b20f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2A/etc/append/crx_append.etc @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/ldrop_append.ma". +include "basic_2/reduction/crx.ma". + +(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************) + +(* Advanved properties ******************************************************) + +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=4 by crx_delta, lt_to_le/ +qed. + +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.