]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma
update in basic_2 ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop_leq.ma
index e5a7555313bbd7f28cc196f846f6664aa79a9cfa..91b20e72029387b813b2bd7d4c04553def7c5d5a 100644 (file)
@@ -12,7 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/ynat/ynat_plus.ma".
 include "basic_2/grammar/leq_leq.ma".
 include "basic_2/relocation/ldrop.ma".
 
@@ -60,6 +59,17 @@ elim (leq_ldrop_trans_be … (leq_sym … HL12) … HLK1) // -L1 -Hdi -Hide
 /3 width=3 by leq_sym, ex2_intro/
 qed-.
 
+lemma ldrop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i →
+                   ∃∃L2. L1 ≃[0, i] L2 & ⇩[i] L2 ≡ K2.
+#K2 #i @(nat_ind_plus … i) -i
+[ /3 width=3 by leq_O2, ex2_intro/
+| #i #IHi #Y #Hi elim (ldrop_O1_lt Y 0) //
+  #I #L1 #V #H lapply (ldrop_inv_O2 … H) -H #H destruct
+  normalize in Hi; elim (IHi L1) -IHi
+  /3 width=5 by ldrop_drop, leq_pair, injective_plus_l, ex2_intro/
+]
+qed-.
+
 lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
 #R #HR #L1 #K1 #s #d #e #HLK1 #K2 #H elim H -K2
 [ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1
@@ -68,3 +78,15 @@ lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R).
   /3 width=6 by leq_trans, step, ex3_intro/
 ]
 qed-.
+
+(* Inversion lemmas on equivalence ******************************************)
+
+lemma ldrop_O1_inj: ∀i,L1,L2,K. ⇩[i] L1 ≡ K → ⇩[i] L2 ≡ K → L1 ≃[i, ∞] L2.
+#i @(nat_ind_plus … i) -i
+[ #L1 #L2 #K #H <(ldrop_inv_O2 … H) -K #H <(ldrop_inv_O2 … H) -L1 //
+| #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 //
+  lapply (ldrop_fwd_length … HLK1)
+  <(ldrop_fwd_length … HLK2) [ /4 width=5 by ldrop_inv_drop1, leq_succ/ ]
+  normalize <plus_n_Sm #H destruct
+]
+qed-.