X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fmultiple%2Ffrees_append.ma;h=ee8324cc9c30d2288eeb1bb6408e5a303cf6e655;hb=658c000ee2ea2da04cf29efc0acdaf16364fbf5e;hp=8f683595597b0cbb9edeafb311494d755fb5c4b5;hpb=1994fe8e6355243652770f53a02db5fdf26915f0;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 8f6835955..ee8324cc9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/multiple/frees_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/frees_append.ma @@ -25,9 +25,10 @@ lemma frees_append: ∀L2,U,l,i. L2 ⊢ i ϵ 𝐅*[l]⦃U⦄ → i ≤ |L2| → #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 -l /3 width=3 by lt_to_le, lt_to_le_to_lt/ +[ -I -L1 -K2 -U -W -l /4 width=3 by ylt_yle_trans, ylt_inv_inj, lt_to_le/ | #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/ + >(minus_plus_m_m (|K2|) 1) >H0 -H0 yminus_SO2 + /3 width=1 by monotonic_yle_minus_dx, yle_pred/ ] qed. @@ -37,13 +38,14 @@ fact frees_inv_append_aux: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ → ∀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/ ] +elim (drop_O1_lt (Ⓕ) L2 j) [2: -Z -Y -L1 -X -U -l /3 width=3 by ylt_yle_trans, ylt_inv_inj/ ] #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 -l /3 width=3 by lt_to_le, lt_to_le_to_lt/ +[ -Z -I -Y -K2 -L1 -X -U -W -l /4 width=3 by ylt_yle_trans, ylt_inv_inj, lt_to_le/ | 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/ + >(minus_plus_m_m (|K2|) 1) >H0 -H0 yminus_SO2 + /3 width=1 by monotonic_yle_minus_dx, yle_pred/ ] qed-.