]> matita.cs.unibo.it Git - helm.git/commitdiff
we now use non-telescopic substitution in parallel reduction, rather
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 22 Aug 2011 12:34:34 +0000 (12:34 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Mon, 22 Aug 2011 12:34:34 +0000 (12:34 +0000)
than the telescopic one. This choice weakens the step of
context-sensitive reduction a bit while maintaining the expected
properties

matita/matita/contribs/lambda-delta/Basic-2/reduction/cpr.ma
matita/matita/contribs/lambda-delta/Basic-2/reduction/lcpr.ma
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_lift.ma
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_split.ma
matita/matita/contribs/lambda-delta/Basic-2/substitution/tps_tps.ma

index 21169192cb74c759affea24e2dddcd59e6ed79e6..f41999e45cf56f4e0772d48eb6e5653fdd110d0e 100644 (file)
@@ -29,7 +29,7 @@ lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2.
 /2/ qed.
 
 lemma cpr_tps: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2.
-/3 width=5/ qed. 
+/3 width=5/ qed.
 
 lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T.
 /2/ qed.
@@ -40,11 +40,9 @@ lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
 qed.
 
-lemma cpr_delta: ∀L,K,V1,V2,V,i.
-                 ↓[0, i] L ≡ K. 𝕓{Abbr} V1 → K ⊢ V1 [0, |L| - i - 1] ≫ V2 →
-                 ↑[0, i + 1] V2 ≡ V → L ⊢ #i ⇒ V.
-#L #K #V1 #V2 #V #i #HLK #HV12 #HV2
-@ex2_1_intro [2: // | skip ] /3 width=8/ (**) (* /4/ is too slow *)
+lemma cpr_delta: ∀L,K,V,W,i.
+                 ↓[0, i] L ≡ K. 𝕓{Abbr} V → ↑[0, i + 1] V ≡ W → L ⊢ #i ⇒ W.
+/3/
 qed.
 
 lemma cpr_cast: ∀L,V,T1,T2.
index 1b19df78fb3721abb7122cf333d337e1b553067c..3098a24561d5887edbbaeb7170feba7c1b86fe11 100644 (file)
@@ -19,7 +19,7 @@ include "Basic-2/reduction/cpr.ma".
 inductive lcpr: lenv → lenv → Prop ≝
 | lcpr_sort: lcpr (⋆) (⋆)
 | lcpr_item: ∀K1,K2,I,V1,V2.
-             lcpr K1 K2 → K1 ⊢ V1 ⇒ V2 → lcpr (K1. 𝕓{I} V1) (K2. 𝕓{I} V2) (*𝕓*)
+             lcpr K1 K2 → K2 ⊢ V1 ⇒ V2 → lcpr (K1. 𝕓{I} V1) (K2. 𝕓{I} V2) (*𝕓*)
 .
 
 interpretation
@@ -29,7 +29,7 @@ interpretation
 (* Basic inversion lemmas ***************************************************)
 
 lemma lcpr_inv_item1_aux: ∀L1,L2. L1 ⊢ ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{I} V1 →
-                          ∃∃K2,V2. K1 ⊢ ⇒ K2 & K1 ⊢ V1 ⇒ V2 & L2 = K2. 𝕓{I} V2.
+                          ∃∃K2,V2. K1 ⊢ ⇒ K2 & K2 ⊢ V1 ⇒ V2 & L2 = K2. 𝕓{I} V2.
 #L1 #L2 * -L1 L2
 [ #K1 #I #V1 #H destruct
 | #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct - K1 I V1 /2 width=5/
@@ -37,5 +37,5 @@ lemma lcpr_inv_item1_aux: ∀L1,L2. L1 ⊢ ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{
 qed.
 
 lemma lcpr_inv_item1: ∀K1,I,V1,L2. K1. 𝕓{I} V1 ⊢ ⇒ L2 →
-                      ∃∃K2,V2. K1 ⊢ ⇒ K2 & K1 ⊢ V1 ⇒ V2 & L2 = K2. 𝕓{I} V2.
+                      ∃∃K2,V2. K1 ⊢ ⇒ K2 & K2 ⊢ V1 ⇒ V2 & L2 = K2. 𝕓{I} V2.
 /2/ qed.
index d2c9e636bb19ac3a5853ab9217d00f71c2de36e4..bc82e4e6d2c870a499c1dcb6ed6bba96e5f36d1a 100644 (file)
 
 include "Basic-2/substitution/drop.ma".
 
-(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
+(* PARALLEL SUBSTITUTION ON TERMS *******************************************)
 
 inductive tps: lenv → term → nat → nat → term → Prop ≝
 | tps_sort : ∀L,k,d,e. tps L (⋆k) d e (⋆k)
 | tps_lref : ∀L,i,d,e. tps L (#i) d e (#i)
-| tps_subst: ∀L,K,V,U1,U2,i,d,e.
-             d ≤ i → i < d + e →
-             ↓[0, i] L ≡ K. 𝕓{Abbr} V → tps K V 0 (d + e - i - 1) U1 →
-             ↑[0, i + 1] U1 ≡ U2 → tps L (#i) d e U2
+| tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e →
+             ↓[0, i] L ≡ K. 𝕓{Abbr} V → ↑[0, i + 1] V ≡ W → tps L (#i) d e W
 | tps_bind : ∀L,I,V1,V2,T1,T2,d,e.
              tps L V1 d e V2 → tps (L. 𝕓{I} V1) T1 (d + 1) e T2 →
              tps L (𝕓{I} V1. T1) d e (𝕓{I} V2. T2)
@@ -31,7 +29,7 @@ inductive tps: lenv → term → nat → nat → term → Prop ≝
              tps L (𝕗{I} V1. T1) d e (𝕗{I} V2. T2)
 .
 
-interpretation "partial telescopic substritution"
+interpretation "parallel substritution (term)"
    'PSubst L T1 d e T2 = (tps L T1 d e T2).
 
 (* Basic properties *********************************************************)
@@ -41,9 +39,8 @@ lemma tps_leq_repl: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≫ T2 →
 #L1 #T1 #T2 #d #e #H elim H -H L1 T1 T2 d e
 [ //
 | //
-| #L1 #K1 #V #V1 #V2 #i #d #e #Hdi #Hide #HLK1 #_ #HV12 #IHV12 #L2 #HL12
-  elim (drop_leq_drop1 … HL12 … HLK1 ? ?) -HL12 HLK1 // #K2 #HK12 #HLK2
-  @tps_subst [4,5,6,8: // |1,2,3: skip | /2/ ] (**) (* /3 width=6/ is too slow *)
+| #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12
+  elim (drop_leq_drop1 … HL12 … HLK1 ? ?) -HL12 HLK1 // /2/
 | /4/
 | /3/
 ]
@@ -60,10 +57,9 @@ lemma tps_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ≫ T2 →
 #L #T1 #T #d1 #e1 #H elim H -L T1 T d1 e1
 [ //
 | //
-| #L #K #V #V1 #V2 #i #d1 #e1 #Hid1 #Hide1 #HLK #_ #HV12 #IHV12 #d2 #e2 #Hd12 #Hde12
+| #L #K #V #W #i #d1 #e1 #Hid1 #Hide1 #HLK #HVW #d2 #e2 #Hd12 #Hde12
   lapply (transitive_le … Hd12 … Hid1) -Hd12 Hid1 #Hid2
-  lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 #Hide2
-  @tps_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2/ ] (**) (* /4 width=6/ is too slow *)
+  lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 /2/
 | /4/
 | /4/
 ]
@@ -74,12 +70,10 @@ lemma tps_weak_top: ∀L,T1,T2,d,e.
 #L #T1 #T #d #e #H elim H -L T1 T d e
 [ //
 | //
-| #L #K #V #V1 #V2 #i #d #e #Hdi #_ #HLK #_ #HV12 #IHV12
+| #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW
   lapply (drop_fwd_drop2_length … HLK) #Hi
   lapply (le_to_lt_to_lt … Hdi Hi) #Hd
-  lapply (plus_minus_m_m_comm (|L|) d ?) [ /2/ ] -Hd #Hd
-  lapply (drop_fwd_O1_length … HLK) normalize #HKL
-  lapply (tps_weak … IHV12 0 (|L| - i - 1) ? ?) -IHV12 // -HKL /2 width=6/
+  lapply (plus_minus_m_m_comm (|L|) d ?) /2/
 | normalize /2/
 | /2/
 ]
@@ -90,20 +84,19 @@ lemma tps_weak_all: ∀L,T1,T2,d,e.
 #L #T1 #T #d #e #HT12
 lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12
 lapply (tps_weak_top … HT12) //
-qed.     
+qed.
 
 (* Basic inversion lemmas ***************************************************)
 
 lemma tps_inv_lref1_aux: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. T1 = #i →
                          T2 = #i ∨ 
-                         ∃∃K,V1,V2,i. d ≤ i & i < d + e &
-                                      ↓[O, i] L ≡ K. 𝕓{Abbr} V1 &
-                                      K ⊢ V1 [O, d + e - i - 1] ≫ V2 &
-                                      ↑[O, i + 1] V2 ≡ T2.
+                         ∃∃K,V,i. d ≤ i & i < d + e &
+                                  ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                                  ↑[O, i + 1] V ≡ T2.
 #L #T1 #T2 #d #e * -L T1 T2 d e
 [ #L #k #d #e #i #H destruct
 | /2/
-| #L #K #V1 #V2 #T2 #i #d #e #Hdi #Hide #HLK #HV12 #HVT2 #j #H destruct -i /3 width=9/
+| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #j #H destruct -i /3 width=7/
 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #i #H destruct
 ]
@@ -111,10 +104,9 @@ qed.
 
 lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i [d, e] ≫ T2 →
                      T2 = #i ∨ 
-                     ∃∃K,V1,V2,i. d ≤ i & i < d + e &
-                                  ↓[O, i] L ≡ K. 𝕓{Abbr} V1 &
-                                  K ⊢ V1 [O, d + e - i - 1] ≫ V2 &
-                                  ↑[O, i + 1] V2 ≡ T2.
+                     ∃∃K,V,i. d ≤ i & i < d + e &
+                              ↓[O, i] L ≡ K. 𝕓{Abbr} V &
+                              ↑[O, i + 1] V ≡ T2.
 /2/ qed.
 
 lemma tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
@@ -125,7 +117,7 @@ lemma tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
 #d #e #L #U1 #U2 * -d e L U1 U2
 [ #L #k #d #e #I #V1 #T1 #H destruct
 | #L #i #d #e #I #V1 #T1 #H destruct
-| #L #K #V #U1 #U2 #i #d #e #_ #_ #_ #_ #_ #I #V1 #T1 #H destruct
+| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct
 | #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/
 | #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
 ]
@@ -144,7 +136,7 @@ lemma tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
 #d #e #L #U1 #U2 * -d e L U1 U2
 [ #L #k #d #e #I #V1 #T1 #H destruct
 | #L #i #d #e #I #V1 #T1 #H destruct
-| #L #K #V #U1 #U2 #i #d #e #_ #_ #_ #_ #_ #I #V1 #T1 #H destruct
+| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct
 | #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct
 | #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/
 ]
index 98c73149d750f105d0ef0b42ac9dad5cd905a7d9..e5ad25ff13ebf98ea6fccfb4df3a3560389e39fc 100644 (file)
@@ -26,16 +26,16 @@ lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 →
                    L ⊢ U1 [dt, et] ≫ U2.
 #K #T1 #T2 #dt #et #H elim H -H K T1 T2 dt et
 [ #K #k #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_
-  lapply (lift_mono … H1 … H2) -H1 H2 #H destruct -U1 //
+  >(lift_mono … H1 … H2) -H1 H2 //
 | #K #i #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_
-  lapply (lift_mono … H1 … H2) -H1 H2 #H destruct -U1 //
-| #K #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HKV #_ #HV12 #IHV12 #L #U1 #U2 #d #e #HLK #H #HVU2 #Hdetd
-  lapply (lt_to_le_to_lt … Hidet … Hdetd) #Hid
+  >(lift_mono … H1 … H2) -H1 H2 //
+| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HVU2 #Hdetd
+  lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid
   lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct -U1;
-  elim (lift_trans_ge … HV12 … HVU2 ?) -HV12 HVU2 V2 // <minus_plus #V2 #HV12 #HVU2
-  elim (drop_trans_le … HLK … HKV ?) -HLK HKV K /2/ #X #HLK #H
-  elim (drop_inv_skip2 … H ?) -H /2/ -Hid #K #W #HKV #HVW #H destruct -X
-  @tps_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2 width=6/ ] (**) (* explicit constructor *)
+  elim (lift_trans_ge … HVW … HVU2 ?) -HVW HVU2 W // <minus_plus #W #HVW #HWU2
+  elim (drop_trans_le … HLK … HKV ?) -HLK HKV K [2: /2/] #X #HLK #H
+  elim (drop_inv_skip2 … H ?) -H [2: /2/] -Hid #K #Y #_ #HVY
+  >(lift_mono … HVY … HVW) -HVY HVW Y #H destruct -X /2/
 | #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd
   elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
   elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2;
@@ -54,16 +54,14 @@ lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 [dt, et] ≫ T2 →
                    L ⊢ U1 [dt + e, et] ≫ U2.
 #K #T1 #T2 #dt #et #H elim H -H K T1 T2 dt et
 [ #K #k #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_
-  lapply (lift_mono … H1 … H2) -H1 H2 #H destruct -U1 //
+  >(lift_mono … H1 … H2) -H1 H2 //
 | #K #i #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_
-  lapply (lift_mono … H1 … H2) -H1 H2 #H destruct -U1 //
-| #K #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HKV #HV1 #HV12 #_ #L #U1 #U2 #d #e #HLK #H #HVU2 #Hddt
-  <(arith_c1x ? ? ? e) in HV1 #HV1 (**) (* explicit athmetical rewrite and ?'s *)
+  >(lift_mono … H1 … H2) -H1 H2 //
+| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt
   lapply (transitive_le … Hddt … Hdti) -Hddt #Hid
   lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct -U1;
-  lapply (lift_trans_be … HV12 … HVU2 ? ?) -HV12 HVU2 V2 /2/ >plus_plus_comm_23 #HV1U2
-  lapply (drop_trans_ge_comm … HLK … HKV ?) -HLK HKV K // -Hid #HLKV
-  @tps_subst [4,5: /2/ |6,7,8: // |1,2,3: skip ] (**) (* /3 width=8/ is too slow *)
+  lapply (lift_trans_be … HVW … HWU2 ? ?) -HVW HWU2 W // [ /2/ ] >plus_plus_comm_23 #HVU2
+  lapply (drop_trans_ge_comm … HLK … HKV ?) -HLK HKV K // -Hid /3/
 | #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt
   elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
   elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2;
@@ -84,12 +82,11 @@ lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
   lapply (lift_inv_sort2 … H) -H #H destruct -T1 /2/
 | #L #i #dt #et #K #d #e #_ #T1 #H #_
   elim (lift_inv_lref2 … H) -H * #Hid #H destruct -T1 /3/
-| #L #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HLKV #_ #HV12 #IHV12 #K #d #e #HLK #T1 #H #Hdetd
-  lapply (lt_to_le_to_lt … Hidet … Hdetd) #Hid
+| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd
+  lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid
   lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct -T1;
-  elim (drop_conf_lt … HLK … HLKV ?) -HLK HLKV L // #L #W #HKL #HKVL #HWV
-  elim (IHV12 … HKVL … HWV ?) -HKVL HWV /2/ -Hdetd #W1 #HW1 #HWV1
-  elim (lift_trans_le … HWV1 … HV12 ?) -HWV1 HV12 V1 // >arith_a2 /3 width=6/
+  elim (drop_conf_lt … HLK … HLKV ?) -HLK HLKV L // #L #U #HKL #_ #HUV
+  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 //
@@ -110,16 +107,15 @@ lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
   lapply (lift_inv_sort2 … H) -H #H destruct -T1 /2/
 | #L #i #dt #et #K #d #e #_ #T1 #H #_
   elim (lift_inv_lref2 … H) -H * #Hid #H destruct -T1 /3/
-| #L #KV #V #V1 #V2 #i #dt #et #Hdti #Hidet #HLKV #HV1 #HV12 #_ #K #d #e #HLK #T1 #H #Hdedt
+| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt  
   lapply (transitive_le … Hdedt … Hdti) #Hdei
   lapply (plus_le_weak … Hdedt) -Hdedt #Hedt
-  lapply (plus_le_weak … Hdei) #Hei
-  <(arith_h1 ? ? ? e ? ?) in HV1 // #HV1
+  lapply (plus_le_weak … Hdei) #Hei  
   lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct -T1;
   lapply (drop_conf_ge … HLK … HLKV ?) -HLK HLKV L // #HKV
-  elim (lift_split … HV12 d (i - e + 1) ? ? ?) -HV12; [2,3,4: normalize /2/ ] -Hdei >arith_e2 // #V0 #HV10 #HV02
+  elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW; [2,3,4: normalize /2/ ] -Hdei >arith_e2 // #V0 #HV10 #HV02
   @ex2_1_intro
-  [2: @tps_subst [4: /2/ |6,7,8: // |1,2,3: skip |5: @arith5 // ]  
+  [2: @tps_subst [3: /2/ |5,6: // |1,2: skip |4: @arith5 // ]
   |1: skip
   | //
   ] (**) (* explicitc constructors *)
@@ -141,7 +137,7 @@ lemma tps_inv_lift1_eq: ∀L,U1,U2,d,e.
 #L #U1 #U2 #d #e #H elim H -H L U1 U2 d e
 [ //
 | //
-| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #_ #_ #_ #_ #T1 #H
+| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #T1 #H
   elim (lift_inv_lref2 … H) -H * #H
   [ lapply (le_to_lt_to_lt … Hdi … H) -Hdi H #H
     elim (lt_refl_false … H)
index 041bb0c9817c8fc217e40bb996f2c40878e53ac0..6bb476fcd6ef44d8da10e4cf4057d5138ee3a624 100644 (file)
@@ -23,19 +23,13 @@ lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. d ≤ i →
 #L #T1 #T2 #d #e #H elim H -L T1 T2 d e
 [ /2/
 | /2/
-| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #HV1 #HV12 #IHV12 #j #Hdj #Hjde
-  elim (lt_or_ge i j) #Hij
-  [ -HV1 Hide;
-    lapply (drop_fwd_drop2 … HLK) #HLK'
-    elim (IHV12 (j - i - 1) ? ?) -IHV12; normalize /2/ -Hjde <minus_n_O >arith_b2 // #W1 #HVW1 #HWV1
-    generalize in match HVW1 generalize in match Hij -HVW1 (**) (* rewriting in the premises, rewrites in the goal too *)
-    >(plus_minus_m_m_comm … Hdj) in ⊢ (% → % → ?) -Hdj #Hij' #HVW1
-    elim (lift_total W1 0 (i + 1)) #W2 #HW12
-    lapply (tps_lift_ge … HWV1 … HLK' HW12 HV12 ?) -HWV1 HLK' HV12 // >arith_a2 /3 width=6/
-  | -IHV12 Hdi Hdj;
-    generalize in match HV1 generalize in match Hide -HV1 Hide (**) (* rewriting in the premises, rewrites in the goal too *)
-    >(plus_minus_m_m_comm … Hjde) in ⊢ (% → % → ?) -Hjde #Hide #HV1
-    @ex2_1_intro [2: @tps_lref |1: skip | /2 width=6/ ] (**) (* /3 width=6 is too slow *)
+| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde
+  elim (lt_or_ge i j)
+  [ -Hide Hjde;
+    >(plus_minus_m_m_comm j d) in ⊢ (% → ?) // -Hdj /3/
+  | -Hdi Hdj; #Hid
+    generalize in match Hide -Hide (**) (* rewriting in the premises, rewrites in the goal too *)
+    >(plus_minus_m_m_comm … Hjde) in ⊢ (% → ?) -Hjde /4/
   ]
 | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
   elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
index 857bc7ebe1cb338cc671ccfa72a75e3cfad11a04..1e3c3629a189e6ba42f96c5f3e130de5e235f5e1 100644 (file)
@@ -36,14 +36,13 @@ theorem tps_trans: ∀L,T1,T,d,e. L ⊢ T1 [d, e] ≫ T → ∀T2. L ⊢ T [d, e
 qed.
 *)
 
-axiom tps_conf_subst_subst_lt: ∀L,K1,V1,W1,T1,i1,d,e,T2,K2,V2,W2,i2.
+axiom tps_conf_subst_subst_lt: ∀L,K1,V1,T1,i1,d,e,T2,K2,V2,i2.
    ↓[O, i1] L ≡ K1. 𝕓{Abbr} V1 → ↓[O, i2] L≡ K2. 𝕓{Abbr} V2 →
-   K1 ⊢ V1 [O, d + e - i1 - 1] ≫ W1 → K2 ⊢ V2 [O, d + e - i2 - 1] ≫ W2 →
-   ↑[O, i1 + 1] W1 ≡ T1 → ↑[O, i2 + 1] W2 ≡ T2 → 
+   ↑[O, i1 + 1] V1 ≡ T1 → ↑[O, i2 + 1] V2 ≡ T2 → 
    d ≤ i1 → i1 < d + e → d ≤ i2 → i2 < d + e → i1 < i2 →
    ∃∃T. L ⊢ T1 [d, e] ≫ T & L ⊢ T2 [d, e] ≫ T.  
 (*
-#L #K1 #V1 #W1 #T1 #i1 #d #e #T2 #K2 #V2 #W2 #i2
+#L #K1 #V1 #T1 #i1 #d #e #T2 #K2 #V2 #i2
 #HLK1 #HLK2 #HVW1 #HVW2 #HWT1 #HWT2 #Hdi1 #Hi1de #Hdi2 #Hi2de #Hi12
 *)
 
@@ -52,25 +51,16 @@ theorem tps_conf: ∀L,T0,T1,d,e. L ⊢ T0 [d, e] ≫ T1 → ∀T2. L ⊢ T0 [d,
 #L #T0 #T1 #d #e #H elim H -H L T0 T1 d e
 [ /2/
 | /2/
-| #L #K1 #V1 #W1 #T1 #i1 #d #e #Hdi1 #Hi1de #HLK1 #HVW1 #HWT1 #IHVW1 #T2 #H
+| #L #K1 #V1 #T1 #i1 #d #e #Hdi1 #Hi1de #HLK1 #HVT1 #T2 #H
   elim (tps_inv_lref1 … H) -H
-  [ -IHVW1 #HX destruct -T2
-    @ex2_1_intro [2: // | skip ] /2 width=6/ (**) (* /3 width=9/ is slow *)
-  | * #K2 #V2 #W2 #i2 #Hdi2 #Hi2de #HLK2 #HVW2 #HWT2
+  [ #HX destruct -T2 /4/
+  | * #K2 #V2 #i2 #Hdi2 #Hi2de #HLK2 #HVT2
     elim (lt_or_eq_or_gt i1 i2) #Hi12
-    [ @tps_conf_subst_subst_lt /width=19/
-    | -HVW1; destruct -i2;
-      lapply (transitive_le ? ? (i1 + 1) Hdi2 ?) -Hdi2 // #Hdi2
-      lapply (drop_mono … HLK1 … HLK2) -HLK1 Hdi1 Hi1de #H destruct -V1 K1;
-      elim (IHVW1 … HVW2) -IHVW1 HVW2 #W #HW1 #HW2
-      lapply (drop_fwd_drop2 … HLK2) -HLK2 #HLK2
-      elim (lift_total W 0 (i1 + 1)) #T #HWT
-      lapply (tps_lift_ge … HW1 … HLK2 HWT1 HWT ?) -HW1 HWT1 //
-      lapply (tps_lift_ge … HW2 … HLK2 HWT2 HWT ?) -HW2 HWT2 HLK2 HWT // normalize #HT2 #HT1
-      lapply (tps_weak … HT1 d e ? ?) -HT1 [ >arith_i2 // | // ]
-      lapply (tps_weak … HT2 d e ? ?) -HT2 [ >arith_i2 // | // ]
-      /2/
-    | @ex2_1_comm @tps_conf_subst_subst_lt /width=19/
+    [ @tps_conf_subst_subst_lt /width=15/
+    | -Hdi2 Hi2de; destruct -i2;
+      lapply (drop_mono … HLK1 … HLK2) -HLK1 #H destruct -V1 K1
+      >(lift_mono … HVT1 … HVT2) -HVT1 /2/
+    | @ex2_1_comm @tps_conf_subst_subst_lt /width=15/
     ]
   ]
 | #L #I #V0 #V1 #T0 #T1 #d #e #_ #_ #IHV01 #IHT01 #X #HX