X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fmultiple%2Ffrees_append.ma;h=8f683595597b0cbb9edeafb311494d755fb5c4b5;hb=c60524dec7ace912c416a90d6b926bee8553250b;hp=e43517ff90c04d9008b96de6eabe6e9ac60e4adb;hpb=f10cfe417b6b8ec1c7ac85c6ecf5fb1b3fdf37db;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/frees_append.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/frees_append.ma index e43517ff9..8f6835955 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/multiple/frees_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/frees_append.ma @@ -19,13 +19,13 @@ include "basic_2/multiple/frees.ma". (* Properties on append for local environments ******************************) -lemma frees_append: ∀L2,U,d,i. L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ → i ≤ |L2| → - ∀L1. L1 @@ L2 ⊢ i ϵ 𝐅*[d]⦃U⦄. -#L2 #U #d #i #H elim H -L2 -U -d -i /3 width=2 by frees_eq/ -#I #L2 #K2 #U #W #d #i #j #Hdj #Hji #HnU #HLK2 #_ #IHW #Hi #L1 +lemma frees_append: ∀L2,U,l,i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → i ≤ |L2| → + ∀L1. L1 @@ L2 ⊢ i ϵ 𝐅*[l]⦃U⦄. +#L2 #U #l #i #H elim H -L2 -U -l -i /3 width=2 by frees_eq/ +#I #L2 #K2 #U #W #l #i #j #Hlj #Hji #HnU #HLK2 #_ #IHW #Hi #L1 lapply (drop_fwd_length_minus2 … HLK2) normalize #H0 lapply (drop_O1_append_sn_le … HLK2 … L1) -HLK2 -[ -I -L1 -K2 -U -W -d /3 width=3 by lt_to_le, lt_to_le_to_lt/ +[ -I -L1 -K2 -U -W -l /3 width=3 by lt_to_le, lt_to_le_to_lt/ | #HLK2 @(frees_be … HnU HLK2) // -HnU -HLK2 @IHW -IHW >(minus_plus_m_m (|K2|) 1) >H0 -H0 /2 width=1 by monotonic_le_minus_l2/ ] @@ -33,20 +33,20 @@ qed. (* Inversion lemmas on append for local environments ************************) -fact frees_inv_append_aux: ∀L,U,d,i. L ⊢ i ϵ 𝐅*[d]⦃U⦄ → ∀L1,L2. L = L1 @@ L2 → - i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄. -#L #U #d #i #H elim H -L -U -d -i /3 width=2 by frees_eq/ -#Z #L #Y #U #X #d #i #j #Hdj #Hji #HnU #HLY #_ #IHW #L1 #L2 #H #Hi destruct -elim (drop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -d /2 width=3 by lt_to_le_to_lt/ ] +fact frees_inv_append_aux: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ → ∀L1,L2. L = L1 @@ L2 → + i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[l]⦃U⦄. +#L #U #l #i #H elim H -L -U -l -i /3 width=2 by frees_eq/ +#Z #L #Y #U #X #l #i #j #Hlj #Hji #HnU #HLY #_ #IHW #L1 #L2 #H #Hi destruct +elim (drop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -l /2 width=3 by lt_to_le_to_lt/ ] #I #K2 #W #HLK2 lapply (drop_fwd_length_minus2 … HLK2) normalize #H0 lapply (drop_O1_inv_append1_le … HLY … HLK2) -HLY -[ -Z -I -Y -K2 -L1 -X -U -W -d /3 width=3 by lt_to_le, lt_to_le_to_lt/ +[ -Z -I -Y -K2 -L1 -X -U -W -l /3 width=3 by lt_to_le, lt_to_le_to_lt/ | normalize #H destruct @(frees_be … HnU HLK2) -HnU -HLK2 // @IHW -IHW // >(minus_plus_m_m (|K2|) 1) >H0 -H0 /2 width=1 by monotonic_le_minus_l2/ ] qed-. -lemma frees_inv_append: ∀L1,L2,U,d,i. L1 @@ L2 ⊢ i ϵ 𝐅*[d]⦃U⦄ → - i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[d]⦃U⦄. +lemma frees_inv_append: ∀L1,L2,U,l,i. L1 @@ L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → + i ≤ |L2| → L2 ⊢ i ϵ 𝐅*[l]⦃U⦄. /2 width=4 by frees_inv_append_aux/ qed-.