]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_append.ma
- we introduced the pointer_step rc in the perspective of proving
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / substitution / ldrop_append.ma
index ec5258552a4a50a34a50681fb7ac52a544828a1c..359d39c80c4ea0a59c82a9809a79504730cb1a29 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/lenv_append.ma".
 include "basic_2/substitution/ldrop.ma".
 
 (* DROPPING *****************************************************************)
 
 (* Properties on append for local environments ******************************)
 
-lemma ldrop_append_O1_ge: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K →
-                          |L2| ≤ e → ⇩[0, e - |L2|] L1 ≡ K.
+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| →
+                             ∀L. ⇩[0, e] L @@ L1 ≡ L @@ L2.
+/2 width=3 by ldrop_O1_append_sn_le_aux/ qed.
+
+(* Inversion lemmas on append for local environments ************************)
+
+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
@@ -28,16 +41,22 @@ elim (ldrop_inv_O1 … H) -H * #H2e #HL12 destruct
   >commutative_plus normalize #H destruct
 | <minus_plus >minus_minus_comm /3 width=1/
 ]
-qed.
-
-lemma ldrop_append_O1_lt: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → e < |L2| →
-                          ∃∃K2. ⇩[0, e] L2 ≡ K2 & K = L1 @@ K2.
-#K #L1 #L2 elim L2 -L2
-[ #e #_ #H elim (lt_zero_false … H)
-| #L2 #I #V #IHL2 #e normalize #HL12 #H1e
-  elim (ldrop_inv_O1 … HL12) -HL12 * #H2e #HL12 destruct
-  [ -H1e -IHL2 /3 width=3/
-  | elim (IHL2 … HL12 ?) -IHL2 -HL12 /2 width=1/ -H1e #K2 #HLK2 #H destruct /3 width=3/
+qed-.
+
+lemma ldrop_O1_inv_append1_le: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → e ≤ |L2| →
+                               ∀K2. ⇩[0, e] L2 ≡ K2 → K = L1 @@ K2.
+#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 //
+| #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 #H1e #K2 #H2
+    lapply (ldrop_inv_ldrop1 … H1 ?) -H1 //
+    lapply (ldrop_inv_ldrop1 … H2 ?) -H2 // /3 width=4/
   ]
 ]
-qed.
+qed-.