]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop.ma
- lambdadelta: first commutation property on lazy equivalence for
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop.ma
index 94c6b0f733c2a3037932f8fdfd2c3acc6de847a7..bce477f2166b02c53479e5f7a33045a92af1ba96 100644 (file)
@@ -15,7 +15,7 @@
 include "ground_2/lstar.ma".
 include "basic_2/notation/relations/rdrop_4.ma".
 include "basic_2/grammar/lenv_length.ma".
-include "basic_2/grammar/lenv_weight.ma".
+include "basic_2/grammar/cl_restricted_weight.ma".
 include "basic_2/relocation/lift.ma".
 
 (* LOCAL ENVIRONMENT SLICING ************************************************)
@@ -290,6 +290,14 @@ lemma ldrop_fwd_length_lt4: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → 0 < e → |L2|
 #L1 #L2 #d #e #H lapply (ldrop_fwd_length … H) -H /2 width=1/
 qed-.
 
+lemma ldrop_fwd_length_eq: ∀L1,L2,K1,K2,d,e. ⇩[d, e] L1 ≡ K1 → ⇩[d, e] L2 ≡ K2 →
+                           |L1| = |L2| → |K1| = |K2|.
+#L1 #L2 #K1 #K2 #d #e #HLK1 #HLK2 #HL12
+lapply (ldrop_fwd_length … HLK1) -HLK1
+lapply (ldrop_fwd_length … HLK2) -HLK2
+/2 width=2 by injective_plus_r/ (**) (* full auto fails *)
+qed-.
+
 lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}.
 #L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize
 [ /2 width=3/
@@ -309,6 +317,11 @@ lemma ldrop_fwd_lw_lt: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → 0 < e → ♯{L2} <
 ]
 qed-.
 
+lemma ldrop_fwd_rfw: ∀I,L,K,V,i. ⇩[O, i] L ≡ K.ⓑ{I}V → ♯{K, V} < ♯{L, #i}.
+#I #L #K #V #i #HLK lapply (ldrop_fwd_lw … HLK) -HLK
+normalize in ⊢ (%→?%%); /2 width=1 by le_S_S/
+qed-.
+
 (* Advanced inversion lemmas ************************************************)
 
 fact ldrop_inv_O2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → e = 0 → L1 = L2.