]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/substitution/lift_lift.ma
- a caracterization of the top elements of the local evironment
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / substitution / lift_lift.ma
index 74edd2ea2df078c4b82a9e55ea87bd6de029dbf6..727d486d4dfd4e978620ef2c530d71db4685b0c6 100644 (file)
@@ -69,7 +69,7 @@ theorem lift_div_le: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T →
 ]
 qed.
 
-(* Note: apparently this was missing in Basic_1 *)
+(* Note: apparently this was missing in basic_1 *)
 theorem lift_div_be: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T →
                      ∀e,e2,T2. ⇧[d1 + e, e2] T2 ≡ T →
                      e ≤ e1 → e1 ≤ e + e2 →
@@ -198,3 +198,12 @@ theorem lift_trans_ge: ∀d1,e1,T1,T. ⇧[d1, e1] T1 ≡ T →
   elim (IHT12 … HT20 ?) -IHT12 -HT20 // /3 width=5/
 ]
 qed.
+
+(* Advanced properties ******************************************************)
+
+lemma lift_conf_le: ∀T,T1,d. ⇧[O, d] T ≡ T1 → ∀T2,e. ⇧[O, d + e] T ≡ T2 →
+                    ⇧[d, e] T1 ≡ T2.
+#T #T1 #d #HT1 #T2 #e #HT2
+elim (lift_split … HT2 d d ? ? ?) -HT2 // #X #H
+>(lift_mono … H … HT1) -T //
+qed.