X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fsubstitution%2Fdrop_append.ma;h=76198fb23a0d09459c653365d1ea7d06498c9b48;hb=aeec9312d6f72526a460518a1e889eac71657cdd;hp=8aa5a2527548dd8609376e3c0fc9b6bd74e85eae;hpb=1fd63df4c77f5c24024769432ea8492748b4ac79;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/substitution/drop_append.ma b/matita/matita/contribs/lambdadelta/basic_2A/substitution/drop_append.ma index 8aa5a2527..76198fb23 100644 --- a/matita/matita/contribs/lambdadelta/basic_2A/substitution/drop_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2A/substitution/drop_append.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "ground_2/lib/arith_2a.ma". +include "ground/lib/arith_2a.ma". include "basic_2A/grammar/lenv_append.ma". include "basic_2A/substitution/drop.ma". @@ -22,19 +22,19 @@ include "basic_2A/substitution/drop.ma". fact drop_O1_append_sn_le_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → l = 0 → m ≤ |L1| → - ∀L. ⬇[s, 0, m] L @@ L1 ≡ L @@ L2. + ∀L. ⬇[s, 0, m] L ● L1 ≡ L ● L2. #L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m normalize [2,3,4: /4 width=1 by drop_skip_lt, drop_drop, arith_b1, lt_minus_to_plus_r, monotonic_pred/ ] #l #m #_ #_ #H <(le_n_O_to_eq … H) -H // qed-. lemma drop_O1_append_sn_le: ∀L1,L2,s,m. ⬇[s, 0, m] L1 ≡ L2 → m ≤ |L1| → - ∀L. ⬇[s, 0, m] L @@ L1 ≡ L @@ L2. + ∀L. ⬇[s, 0, m] L ● L1 ≡ L ● L2. /2 width=3 by drop_O1_append_sn_le_aux/ qed. (* Inversion lemmas on append for local environments ************************) -lemma drop_O1_inv_append1_ge: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 @@ L2 ≡ K → +lemma drop_O1_inv_append1_ge: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 ● L2 ≡ K → |L2| ≤ m → ⬇[s, 0, m - |L2|] L1 ≡ K. #K #L1 #L2 elim L2 -L2 normalize // #L2 #I #V #IHL2 #s #m #H #H1m @@ -45,8 +45,8 @@ elim (drop_inv_O1_pair1 … H) -H * #H2m #HL12 destruct ] qed-. -lemma drop_O1_inv_append1_le: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 @@ L2 ≡ K → m ≤ |L2| → - ∀K2. ⬇[s, 0, m] L2 ≡ K2 → K = L1 @@ K2. +lemma drop_O1_inv_append1_le: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 ● L2 ≡ K → m ≤ |L2| → + ∀K2. ⬇[s, 0, m] L2 ≡ K2 → K = L1 ● K2. #K #L1 #L2 elim L2 -L2 normalize [ #s #m #H1 #H2 #K2 #H3 lapply (le_n_O_to_eq … H2) -H2 #H2 elim (drop_inv_atom1 … H3) -H3 #H3 #_ destruct