X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fllpx_sn_frees.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fllpx_sn_frees.ma;h=81eff3bda63afeb13d2b2d984403dec7badcf0e8;hb=7593c0f74b944fb100493fb24b665ce3b8d1d252;hp=0000000000000000000000000000000000000000;hpb=361a91ade954f92013da892c62d41e3a7168cfc0;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/llpx_sn_frees.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/llpx_sn_frees.ma new file mode 100644 index 000000000..81eff3bda --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/llpx_sn_frees.ma @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||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/multiple/frees.ma". +include "basic_2/multiple/llpx_sn_alt_rec.ma". + +(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****) + +(* Properties on context-sensitive free variables ***************************) + +fact llpx_sn_frees_trans_aux: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) → + ∀L2,U,l,i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → + ∀L1. llpx_sn R l U L1 L2 → L1 ⊢ i ϵ 𝐅*[l]⦃U⦄. +#R #H1R #H2R #L2 #U #l #i #H elim H -L2 -U -l -i /3 width=2 by frees_eq/ +#I2 #L2 #K2 #U #W2 #l #i #j #Hlj #Hji #HnU #HLK2 #_ #IHW2 #L1 #HL12 +elim (llpx_sn_inv_alt_r … HL12) -HL12 #HL12 #IH +lapply (drop_fwd_length_lt2 … HLK2) #Hj +elim (drop_O1_lt (Ⓕ) L1 j) // -Hj -HL12 #I1 #K1 #W1 #HLK1 +elim (IH … HnU HLK1 HLK2) // -IH -HLK2 /5 width=11 by frees_be/ +qed-. + +lemma llpx_sn_frees_trans: ∀R. (s_r_confluent1 … R (llpx_sn R 0)) → (frees_trans R) → + ∀L1,L2,U,l. llpx_sn R l U L1 L2 → + ∀i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → L1 ⊢ i ϵ 𝐅*[l]⦃U⦄. +/2 width=6 by llpx_sn_frees_trans_aux/ qed-.