]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/substitution/lift_lift.ma
- star.ma: constructor inj of star conflicts with previous constructor
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / substitution / lift_lift.ma
index 74edd2ea2df078c4b82a9e55ea87bd6de029dbf6..f302b16ad4146fa53407b466ce34597cd21746d3 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,20 @@ 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_O1: ∀T,T1,d1,e1. ⇧[d1, e1] T ≡ T1 → ∀T2,e2. ⇧[0, e2] T ≡ T2 →
+                    ∃∃T0. ⇧[0, e2] T1 ≡ T0 & ⇧[d1 + e2, e1] T2 ≡ T0.
+#T #T1 #d1 #e1 #HT1 #T2 #e2 #HT2
+elim (lift_total T1 0 e2) #T0 #HT10
+elim (lift_trans_le … HT1 … HT10 ?) -HT1 // #X #HTX #HT20
+lapply (lift_mono … HTX … HT2) -T #H destruct /2 width=3/
+qed.
+
+lemma lift_conf_be: ∀T,T1,d,e1. ⇧[d, e1] T ≡ T1 → ∀T2,e2. ⇧[d, e2] T ≡ T2 →
+                    e1 ≤ e2 → ⇧[d + e1, e2 - e1] T1 ≡ T2.
+#T #T1 #d #e1 #HT1 #T2 #e2 #HT2 #He12
+elim (lift_split … HT2 (d+e1) e1 ? ? ?) -HT2 // #X #H
+>(lift_mono … H … HT1) -T //
+qed.