]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_append.ma
update in ground_2, web page for ground_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop_append.ma
index 7629115761cdb8f186a349817d90775b68f92468..9b49fdd5bf875116b56b912af27c8b4bf9dab4a4 100644 (file)
@@ -12,6 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "basic_2/grammar/lenv_append.ma".
 include "basic_2/relocation/ldrop.ma".
 
 (* DROPPING *****************************************************************)
@@ -22,8 +23,6 @@ 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 //
 qed-.
 
 lemma ldrop_O1_append_sn_le: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → e ≤ |L1| →
@@ -36,7 +35,7 @@ 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_plus >minus_minus_comm /3 width=1/
@@ -48,12 +47,12 @@ 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 //
+    lapply (ldrop_inv_O2 … H1) -H1 #H1
+    lapply (ldrop_inv_O2 … H2) -H2 #H2 destruct //
   | #e #_ #H1 #H #K2 #H2
     lapply (le_plus_to_le_r … H) -H
     lapply (ldrop_inv_ldrop1 … H1 ?) -H1 //