]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma
unfold on terms completed!
[helm.git] / matita / matita / contribs / lambda-delta / Basic-2 / substitution / tps_lift.ma
index d1d2bbb24ea1f49f2b678f0b2d937134844daa99..c5fca2ca3ce404a45c1f6e2f17fce5fd1d30684b 100644 (file)
@@ -17,8 +17,29 @@ include "Basic-2/substitution/tps.ma".
 
 (* PARTIAL SUBSTITUTION ON TERMS ********************************************)
 
+(* Advanced inversion lemmas ************************************************)
+
+fact tps_inv_refl_SO2_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → e = 1 →
+                           ∀K,V. ↓[0, d] L ≡ K. 𝕓{Abst} V → T1 = T2.
+#L #T1 #T2 #d #e #H elim H -H L T1 T2 d e
+[ //
+| #L #K0 #V0 #W #i #d #e #Hdi #Hide #HLK0 #_ #H destruct -e;
+  >(le_to_le_to_eq … Hdi ?) /2/ -d #K #V #HLK
+  lapply (drop_mono … HLK0 … HLK) #H destruct
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK
+  >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 K V) -IHT12 /2/
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK
+  >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 … HLK) -IHT12 //
+]
+qed.
+
+lemma tps_inv_refl_SO2: ∀L,T1,T2,d. L ⊢ T1 [d, 1] ≫ T2 →
+                        ∀K,V. ↓[0, d] L ≡ K. 𝕓{Abst} V → T1 = T2.
+/2 width=8/ qed.
+
 (* Relocation properties ****************************************************)
 
+(* Basic-1: was: subst1_lift_lt *)
 lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 →
                    ∀L,U1,U2,d,e. ↓[d, e] L ≡ K →
                    ↑[d, e] T1 ≡ U1 → ↑[d, e] T2 ≡ U2 →
@@ -45,6 +66,7 @@ lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 →
 ]
 qed.
 
+(* Basic-1: was: subst1_lift_ge *)
 lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 →
                    ∀L,U1,U2,d,e. ↓[d, e] L ≡ K →
                    ↑[d, e] T1 ≡ U1 → ↑[d, e] T2 ≡ U2 →
@@ -69,6 +91,7 @@ lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 →
 ]
 qed.
 
+(* Basic-1: was: subst1_gen_lift_lt *)
 lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
                         ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 →
                         dt + et ≤ d →
@@ -85,8 +108,9 @@ lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
   elim (lift_trans_le … HUV … HVW ?) -HUV HVW V // >arith_a2 // -Hid /3/
 | #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
   elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X;
-  elim (IHV12 … HLK … HWV1 ?) -IHV12 //
-  elim (IHU12 … HTU1 ?) -IHU12 HTU1 [3: /2/ |4: @drop_skip // |2: skip ] -HLK HWV1 Hdetd /3 width=5/ (**) (* just /3 width=5/ is too slow *)
+  elim (IHV12 … HLK … HWV1 ?) -IHV12 HWV1 // #W2 #HW12 #HWV2
+  elim (IHU12 … HTU1 ?) -IHU12 HTU1 [3: /2/ |4: @drop_skip // |2: skip ] -HLK Hdetd (**) (* /3 width=5/ is too slow *)
+  /3 width=5/
 | #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
   elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X;
   elim (IHV12 … HLK … HWV1 ?) -IHV12 HWV1 //
@@ -94,6 +118,7 @@ lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
 ]
 qed.
 
+(* Basic-1: was: subst1_gen_lift_ge *)
 lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
                         ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 →
                         d + e ≤ dt →
@@ -118,7 +143,7 @@ lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
 | #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
   elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X;
   lapply (plus_le_weak … Hdetd) #Hedt
-  elim (IHV12 … HLK … HWV1 ?) -IHV12 // #W2 #HW12 #HWV2
+  elim (IHV12 … HLK … HWV1 ?) -IHV12 HWV1 // #W2 #HW12 #HWV2
   elim (IHU12 … HTU1 ?) -IHU12 HTU1 [4: @drop_skip // |2: skip |3: /2/ ]
   <plus_minus // /3 width=5/
 | #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd
@@ -128,6 +153,7 @@ lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
 ]
 qed.
 
+(* Basic-1: was: subst1_gen_lift_eq *)
 lemma tps_inv_lift1_eq: ∀L,U1,U2,d,e.
                         L ⊢ U1 [d, e] ≫ U2 → ∀T1. ↑[d, e] T1 ≡ U1 → U1 = U2.
 #L #U1 #U2 #d #e #H elim H -H L U1 U2 d e
@@ -148,10 +174,6 @@ lemma tps_inv_lift1_eq: ∀L,U1,U2,d,e.
 ]
 qed.
 (*
-      Theorem subst0_gen_lift_ge : (u,t1,x:?; i,h,d:?) (subst0 i u (lift h d t1) x) ->
-                                   (le (plus d h) i) ->
-                                   (EX t2 | x = (lift h d t2) & (subst0 (minus i h) u t1 t2)).
-
       Theorem subst0_gen_lift_rev_ge: (t1,v,u2:?; i,h,d:?) 
                                       (subst0 i v t1 (lift h d u2)) ->
                                       (le (plus d h) i) ->
@@ -176,23 +198,3 @@ lapply (tps_weak … HU1 d e ? ?) -HU1 // <plus_minus_m_m_comm // -Hddt Hdtde #H
 lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -U1;
 elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 HLK HTU1 // <minus_plus_m_m /2/
 qed.
-
-(* Advanced inversion lemmas ************************************************)
-
-fact tps_inv_refl1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → e = 1 →
-                        ∀K,V. ↓[0, d] L ≡ K. 𝕓{Abst} V → T1 = T2.
-#L #T1 #T2 #d #e #H elim H -H L T1 T2 d e
-[ //
-| #L #K0 #V0 #W #i #d #e #Hdi #Hide #HLK0 #_ #H destruct -e;
-  >(le_to_le_to_eq … Hdi ?) /2/ -d #K #V #HLK
-  lapply (drop_mono … HLK0 … HLK) #H destruct
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK
-  >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 K V) -IHT12 /2/
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H1 #K #V #HLK
-  >(IHV12 H1 … HLK) -IHV12 >(IHT12 H1 … HLK) -IHT12 //
-]
-qed.
-
-lemma tps_inv_refl1: ∀L,T1,T2,d. L ⊢ T1 [d, 1] ≫ T2 →
-                     ∀K,V. ↓[0, d] L ≡ K. 𝕓{Abst} V → T1 = T2.
-/2 width=8/ qed.