X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fldrop_append.ma;h=2da9f9ed9d5f8da316ccb9f13b37b4cb7bcd24a4;hb=f663836446db6814e2d0cbb11638070ede01bdb4;hp=6cae606c7208e2c7ebf47322a2fdfa6bde240227;hpb=e2fd96302d52266bec42a19f100dadc6111fc07b;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma index 6cae606c7..2da9f9ed9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma @@ -22,9 +22,8 @@ include "basic_2/relocation/ldrop.ma". fact ldrop_O1_append_sn_le_aux: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → d = 0 → e ≤ |L1| → ∀L. ⇩[0, e] L @@ L1 ≡ L @@ L2. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize // /4 width=1/ -#d #e #_ #H #L -d -lapply (le_n_O_to_eq … H) -H // +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize +/4 width=1 by ldrop_skip_lt, ldrop_ldrop, arith_b1, lt_minus_to_plus_r, monotonic_pred/ qed-. lemma ldrop_O1_append_sn_le: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → e ≤ |L1| → @@ -37,10 +36,10 @@ lemma ldrop_O1_inv_append1_ge: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → |L2| ≤ e → ⇩[0, e - |L2|] L1 ≡ K. #K #L1 #L2 elim L2 -L2 normalize // #L2 #I #V #IHL2 #e #H #H1e -elim (ldrop_inv_O1 … H) -H * #H2e #HL12 destruct +elim (ldrop_inv_O1_pair1 … H) -H * #H2e #HL12 destruct [ lapply (le_n_O_to_eq … H1e) -H1e -IHL2 >commutative_plus normalize #H destruct -| minus_minus_comm /3 width=1/ +| minus_minus_comm /3 width=1 by monotonic_pred/ ] qed-. @@ -49,17 +48,13 @@ lemma ldrop_O1_inv_append1_le: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → e ≤ #K #L1 #L2 elim L2 -L2 normalize [ #e #H1 #H2 #K2 #H3 lapply (le_n_O_to_eq … H2) -H2 #H2 - lapply (ldrop_inv_atom1 … H3) -H3 #H3 destruct - >(ldrop_inv_refl … H1) -H1 // + elim (ldrop_inv_atom1 … H3) -H3 #H3 #_ destruct + >(ldrop_inv_O2 … H1) -H1 // | #L2 #I #V #IHL2 #e @(nat_ind_plus … e) -e [ -IHL2 ] [ #H1 #_ #K2 #H2 - lapply (ldrop_inv_refl … H1) -H1 #H1 - lapply (ldrop_inv_refl … H2) -H2 #H2 destruct // - | #e #_ #H1 #H #K2 #H2 - lapply (le_plus_to_le_r … H) -H - lapply (ldrop_inv_ldrop1 … H1 ?) -H1 // - lapply (ldrop_inv_ldrop1 … H2 ?) -H2 // -