\ /
V_______________________________________________________________ *)
-include "basics/list.ma".
+include "arithmetics/nat.ma".
(* ARITHMETICAL PROPERTIES **************************************************)
-lemma plus_plus_comm_23: ∀m,n,p. m + n + p = m + p + n.
-// qed.
-
-lemma minus_plus_comm: ∀a,b,c. a - b - c = a - (c + b).
-// qed.
-
-lemma arith8: ∀a,b. a < a + b + 1.
-// qed.
+lemma plus_S_eq_O_false: ∀n,m. n + S m = 0 → False.
+#n #m <plus_n_Sm #H destruct
+qed.
-lemma arith9: ∀a,b,c. c < a + (b + c + 1) + 1.
-// qed.
+lemma plus_S_le_to_pos: ∀n,m,p. n + S m ≤ p → 0 < p.
+#n #m #p <plus_n_Sm #H @(lt_to_le_to_lt … H) //
+qed.
lemma minus_le: ∀m,n. m - n ≤ m.
/2/ qed.
-lemma plus_plus_minus_m_m: ∀e1,e2,d. e1 ≤ e2 → d + e1 + (e2 - e1) = d + e2.
+lemma le_O_to_eq_O: ∀n. n ≤ 0 → n = 0.
/2/ qed.
-lemma le_O_to_eq_O: ∀n. n ≤ 0 → n = 0.
+lemma lt_to_le: ∀a,b. a < b → a ≤ b.
/2/ qed.
+lemma lt_refl_false: ∀n. n < n → False.
+#n #H elim (lt_to_not_eq … H) -H /2/
+qed.
+
+lemma lt_zero_false: ∀n. n < 0 → False.
+#n #H elim (lt_to_not_le … H) -H /2/
+qed.
+
+lemma lt_or_ge: ∀m,n. m < n ∨ n ≤ m.
+#m #n elim (decidable_lt m n) /3/
+qed.
+
+lemma le_to_lt_or_eq: ∀m,n. m ≤ n → m < n ∨ m = n.
+#m #n * -n /3/
+qed.
+
+lemma plus_le_weak: ∀m,n,p. m + n ≤ p → n ≤ p.
+/2/ qed.
+
+lemma plus_lt_false: ∀m,n. m + n < m → False.
+#m #n #H1 lapply (le_plus_n_r n m) #H2
+lapply (le_to_lt_to_lt … H2 H1) -H2 H1 #H
+elim (lt_refl_false … H)
+qed.
+
+lemma monotonic_lt_minus_l: ∀p,q,n. n ≤ q → q < p → q - n < p - n.
+#p #q #n #H1 #H2
+@lt_plus_to_minus_r <plus_minus_m_m //.
+qed.
+
lemma plus_le_minus: ∀a,b,c. a + b ≤ c → a ≤ c - b.
/2/ qed.
-lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → (m + n) - p = (m - p) + n.
+lemma lt_plus_minus: ∀i,u,d. u ≤ i → i < d + u → i - u < d.
+/2/ qed.
+
+lemma plus_plus_comm_23: ∀m,n,p. m + n + p = m + p + n.
+// qed.
+
+lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n.
#n #m #p #lepm @plus_to_minus <associative_plus
>(commutative_plus p) <plus_minus_m_m //
qed.
-lemma le_plus_minus: ∀a,b,c. c ≤ b → a + b - c = a + (b - c).
-/2/ qed.
-
lemma minus_le_minus_minus_comm: ∀m,p,n.
p ≤ m → m - p ≤ n → n + p - m = n - (m - p).
#m elim m -m
]
qed.
-lemma lt_refl_false: ∀n. n < n → False.
-#n #H elim (lt_to_not_eq … H) -H /2/
-qed.
+lemma minus_plus_comm: ∀a,b,c. a - b - c = a - (c + b).
+// qed.
-lemma lt_zero_false: ∀n. n < 0 → False.
-#n #H elim (lt_to_not_le … H) -H /2/
+lemma minus_minus_comm: ∀a,b,c. a - b - c = a - c - b.
+/3/ qed.
+
+lemma le_plus_minus: ∀a,b,c. c ≤ b → a + b - c = a + (b - c).
+/2/ qed.
+
+lemma plus_minus_m_m_comm: ∀n,m. m ≤ n → n = m + (n - m).
+/2/ qed.
+
+theorem minus_plus_m_m_comm: ∀n,m. n = (m + n) - m.
+/2/ qed.
+
+lemma arith_a2: ∀a,c1,c2. c1 + c2 ≤ a → a - c1 - c2 + (c1 + c2) = a.
+/2/ qed.
+
+lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b.
+#a #b #c1 #H >minus_plus @eq_f2 /2/
qed.
-lemma lt_or_ge: ∀m,n. m < n ∨ n ≤ m.
-#m #n elim (decidable_lt m n) /3/
+lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b.
+#a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2/
qed.
-lemma plus_lt_false: ∀m,n. m + n < m → False.
-#m #n #H1 lapply (le_plus_n_r n m) #H2
-lapply (le_to_lt_to_lt … H2 H1) -H2 H1 #H
-elim (lt_refl_false … H)
+lemma arith_c1: ∀a,b,c1. a + c1 - (b + c1) = a - b.
+// qed.
+
+lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b.
+#x #a #b #c1 >plus_plus_comm_23 //
qed.
-lemma plus_S_eq_O_false: ∀n,m. n + S m = 0 → False.
-#n #m <plus_n_Sm #H destruct
-qed.
+lemma arith_d1: ∀a,b,c1. c1 ≤ b → a + c1 + (b - c1) = a + b.
+/2/ qed.
-lemma minus_minus_comm: ∀a,b,c. a - b - c = a - c - b.
+lemma arith_e2: ∀a,c1,c2. a ≤ c1 → c1 + c2 - (c1 - a + c2) = a.
/3/ qed.
-lemma arith1: ∀n,h,m,p. n + h + m ≤ p + h → n + m ≤ p.
+lemma arith_f1: ∀a,b,c1. a + b ≤ c1 → c1 - (c1 - a - b) = a + b.
+#a #b #c1 #H >minus_plus <minus_minus //
+qed.
+
+lemma arith_g1: ∀a,b,c1. c1 ≤ b → a - (b - c1) - c1 = a - b.
/2/ qed.
-lemma arith6: ∀m,n. m < n → n - (n - m - 1) = m + 1.
-#m #n #H >minus_plus <minus_minus //
+lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b →
+ a1 - c1 + a2 - (b - c1) = a1 + a2 - b.
+#a1 #a2 #b #c1 #H1 #H2 <le_plus_minus_comm /2/
qed.
-lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2.
+lemma arith_z1: ∀a,b,c1. a + c1 - b - c1 = a - b.
+// qed.
+
+(* unstable *****************************************************************)
+
+lemma arith1: ∀n,h,m,p. n + h + m ≤ p + h → n + m ≤ p.
/2/ qed.
-lemma arith5: ∀i,h,d. i + h ≤ d → d - i - h + (i + h) = d.
+lemma arith2: ∀j,i,e,d. d + e ≤ i → d ≤ i - e + j.
+#j #i #e #d #H lapply (plus_le_minus … H) -H /2/
+qed.
+
+lemma arith3: ∀a1,a2,b,c1. a1 + a2 ≤ b → a1 + c1 + a2 ≤ b + c1.
/2/ qed.
-lemma arith7: ∀i,d. i ≤ d → d - i + i = d.
+lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2.
/2/ qed.
-lemma arith2: ∀j,i,e,d. d + e ≤ i → d ≤ i - e + j.
-/3/ qed.
+lemma arith5: ∀a,b1,b2,c1. c1 ≤ b1 → c1 ≤ a → a < b1 + b2 → a - c1 < b1 - c1 + b2.
+#a #b1 #b2 #c1 #H1 #H2 #H3
+<le_plus_minus_comm // @monotonic_lt_minus_l //
+qed.
-lemma arith3: ∀m,n,p. p ≤ m → m + n - (m - p + n) = p.
-/3/ qed.
+lemma arith8: ∀a,b. a < a + b + 1.
+// qed.
+
+lemma arith9: ∀a,b,c. c < a + (b + c + 1) + 1.
+// qed.
+lemma arith10: ∀a,b,c,d,e. a ≤ b → c + (a - d - e) ≤ c + (b - d - e).
+#a #b #c #d #e #H
+>minus_plus >minus_plus @monotonic_le_plus_r @monotonic_le_minus_l //
+qed.
(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************)
-(* language *****************************************************************)
+(* *****************************************************************)
notation "hvbox( ⋆ )"
non associative with precedence 90
(* substitution *************************************************************)
+notation "hvbox( T1 break [ d , break e ] ≈ break T2 )"
+ non associative with precedence 45
+ for @{ 'Eq $T1 $d $e $T2 }.
+
notation "hvbox( ↑ [ d , break e ] break T1 ≡ break T2 )"
non associative with precedence 45
for @{ 'RLift $T1 $d $e $T2 }.
non associative with precedence 45
for @{ 'RDrop $L1 $d $e $L2 }.
-notation "hvbox( L ⊢ break (term 90 T1) break [ d , break e ] ≫ break T2 )"
+notation "hvbox( L ⊢ break term 90 T1 break [ d , break e ] ≫ break T2 )"
non associative with precedence 45
for @{ 'PSubst $L $T1 $d $e $T2 }.
\ /
V_______________________________________________________________ *)
-include "lambda-delta/substitution/ps_defs.ma".
+include "lambda-delta/substitution/pts_defs.ma".
(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
(**************************************************************************)
(*
-include "lambda-delta/substitution/lift_fun.ma".
include "lambda-delta/substitution/lift_main.ma".
include "lambda-delta/substitution/drop_main.ma".
*)
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "lambda-delta/reduction/lpr_defs.ma".
-
-(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
-
-axiom tpr_ps_lpr: ∀L1,L2. L1 ⇒ L2 → ∀T1,T2. T1 ⇒ T2 →
- ∀d,e,U1. L1 ⊢ T1 [d, e] ≫ U1 →
- ∃∃U2. U1 ⇒ U2 & L2 ⊢ T2 [d, e] ≫ U2.
-
-lemma tpr_ps_bind: ∀I,V1,V2,T1,T2,U1. V1 ⇒ V2 → T1 ⇒ T2 →
- ⋆. 𝕓{I} V1 ⊢ T1 [0, 1] ≫ U1 →
- ∃∃U2. U1 ⇒ U2 & ⋆. 𝕓{I} V2 ⊢ T2 [0, 1] ≫ U2.
-/3 width=5/ qed.
--- /dev/null
+(**************************************************************************)
+(* ___ *)
+(* ||M|| *)
+(* ||A|| A project by Andrea Asperti *)
+(* ||T|| *)
+(* ||I|| Developers: *)
+(* ||T|| The HELM team. *)
+(* ||A|| http://helm.cs.unibo.it *)
+(* \ / *)
+(* \ / This file is distributed under the terms of the *)
+(* v GNU General Public License Version 2 *)
+(* *)
+(**************************************************************************)
+
+include "lambda-delta/reduction/lpr_defs.ma".
+
+(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
+
+axiom tpr_pts_lpr: ∀L1,L2. L1 ⇒ L2 → ∀T1,T2. T1 ⇒ T2 →
+ ∀d,e,U1. L1 ⊢ T1 [d, e] ≫ U1 →
+ ∃∃U2. U1 ⇒ U2 & L2 ⊢ T2 [d, e] ≫ U2.
+
+lemma tpr_pts_bind: ∀I,V1,V2,T1,T2,U1. V1 ⇒ V2 → T1 ⇒ T2 →
+ ⋆. 𝕓{I} V1 ⊢ T1 [0, 1] ≫ U1 →
+ ∃∃U2. U1 ⇒ U2 & ⋆. 𝕓{I} V2 ⊢ T2 [0, 1] ≫ U2.
+/3 width=5/ qed.
\ /
V_______________________________________________________________ *)
-include "lambda-delta/substitution/lift_fun.ma".
include "lambda-delta/substitution/lift_weight.ma".
-include "lambda-delta/substitution/ps_ps.ma".
+include "lambda-delta/substitution/pts_pts.ma".
include "lambda-delta/reduction/tpr_main.ma".
-include "lambda-delta/reduction/tpr_ps.ma".
-
-lemma ps_inv_lift_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
-[ //
-| //
-| #L #K #V #U1 #U2 #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)
- | lapply (lt_to_le_to_lt … Hide … H) -Hide H #H
- elim (lt_refl_false … H)
- ]
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
- elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #H destruct -X
- >IHV12 // >IHT12 //
-| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
- elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct -X
- >IHV12 // >IHT12 //
-]
-qed.
+include "lambda-delta/reduction/tpr_pts.ma".
(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
#V0 #V1 #T0 #T1 #V2 #T2 #T #IH #HV01 #HV02 #HT01 #HT02 #HT2
elim (IH … HV01 … HV02) -HV01 HV02 // #V #HV1 #HV2
elim (IH … HT01 … HT02) -HT01 HT02 IH // -V0 T0 #T0 #HT10 #HT20
-elim (tpr_ps_bind … HV2 HT20 … HT2) -HT20 HT2 /3 width=5/
+elim (tpr_pts_bind … HV2 HT20 … HT2) -HT20 HT2 /3 width=5/
qed.
lemma tpr_conf_bind_zeta:
| -HV2 HVV2 #WW2 #UU2 #UU #HWW2 #HUU02 #HUU2 #H destruct -T1;
elim (IH … HW02 … HWW2) -HW02 HWW2 // #W #HW02 #HWW2
elim (IH … HU02 … HUU02) -HU02 HUU02 IH // #U #HU2 #HUUU2
- elim (tpr_ps_bind … HWW2 HUUU2 … HUU2) -HUU2 HUUU2 #UUU #HUUU2 #HUUU1
+ elim (tpr_pts_bind … HWW2 HUUU2 … HUU2) -HUU2 HUUU2 #UUU #HUUU2 #HUUU1
@ex2_1_intro
- [2: @tpr_theta
+ [2: @tpr_theta [6: @HVV |7: @HWW2 |8: @HUUU2 |1,2,3,4: skip | // ]
|1:skip
- |3: @tpr_delta [3: @tpr_flat |1: skip ]
- ] /2 width=14/ (**) (* /5 width=14/ is too slow *)
+ |3: @tpr_delta [3: @tpr_flat |1: skip ] /2 width=5/
+ ] (**) (* /5 width=14/ is too slow *)
(* case 3: zeta *)
| -HW02 HVV HVVV #UU1 #HUU10 #HUUT1
elim (tpr_inv_lift … HU02 … HUU10) -HU02 #UU #HUU2 #HUU1
#V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2
elim (IH … HV01 … HV02) -HV01 HV02 // #V #HV1 #HV2
elim (IH … HT01 … HT02) -HT01 HT02 IH // #T #HT1 #HT2
-elim (tpr_ps_bind … HV1 HT1 … HTT1) -HT1 HTT1 #U1 #TTU1 #HTU1
-elim (tpr_ps_bind … HV2 HT2 … HTT2) -HT2 HTT2 #U2 #TTU2 #HTU2
-elim (ps_conf … HTU1 … HTU2) -HTU1 HTU2 #U #HU1 #HU2
+elim (tpr_pts_bind … HV1 HT1 … HTT1) -HT1 HTT1 #U1 #TTU1 #HTU1
+elim (tpr_pts_bind … HV2 HT2 … HTT2) -HT2 HTT2 #U2 #TTU2 #HTU2
+elim (pts_conf … HTU1 … HTU2) -HTU1 HTU2 #U #HU1 #HU2
@ex2_1_intro [2,3: @tpr_delta |1: skip ] /width=10/ (**) (* /3 width=10/ is too slow *)
qed.
∃∃X. 𝕓{Abbr} V1. TT1 ⇒ X & X2 ⇒ X.
#X2 #V0 #V1 #T0 #T1 #TT1 #T2 #IH #_ #HT01 #HTT1 #HTX2 #HTT20
elim (tpr_inv_lift … HT01 … HTT20) -HT01 #TT2 #HTT21 #HTT2
-lapply (ps_inv_lift_eq … HTT1 … HTT21) -HTT1 #HTT1 destruct -T1;
+lapply (pts_inv_lift1_eq … HTT1 … HTT21) -HTT1 #HTT1 destruct -T1;
lapply (tw_lift … HTT20) -HTT20 #HTT20
elim (IH … HTX2 … HTT2) -HTX2 HTT2 IH /3/
qed.
\ /
V_______________________________________________________________ *)
-include "lambda-delta/syntax/lenv.ma".
+include "lambda-delta/substitution/leq_defs.ma".
include "lambda-delta/substitution/lift_defs.ma".
(* DROPPING *****************************************************************)
interpretation "dropping" 'RDrop L1 d e L2 = (drop L1 d e L2).
-(* Basic properties *********************************************************)
-
-lemma drop_drop_lt: ∀L1,L2,I,V,e.
- ↓[0, e - 1] L1 ≡ L2 → 0 < e → ↓[0, e] L1. 𝕓{I} V ≡ L2.
-#L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) /2/
-qed.
-
(* Basic inversion lemmas ***************************************************)
lemma drop_inv_refl_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2.
-#d #e #L1 #L2 #H elim H -H d e L1 L2
+#d #e #L1 #L2 * -d e L1 L2
[ //
-| #L1 #L2 #I #V #e #_ #_ #_ #H
+| #L1 #L2 #I #V #e #_ #_ #H
elim (plus_S_eq_O_false … H)
-| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #_ #H
+| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H
elim (plus_S_eq_O_false … H)
]
qed.
lemma drop_inv_refl: ∀L1,L2. ↓[0, 0] L1 ≡ L2 → L1 = L2.
/2 width=5/ qed.
+lemma drop_inv_sort1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → L1 = ⋆ →
+ ∧∧ L2 = ⋆ & d = 0 & e = 0.
+#d #e #L1 #L2 * -d e L1 L2
+[ /2/
+| #L1 #L2 #I #V #e #_ #H destruct
+| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct
+]
+qed.
+
+lemma drop_inv_sort1: ∀d,e,L2. ↓[d, e] ⋆ ≡ L2 →
+ ∧∧ L2 = ⋆ & d = 0 & e = 0.
+/2/ qed.
+
lemma drop_inv_O1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 →
∀K,I,V. L1 = K. 𝕓{I} V →
(e = 0 ∧ L2 = K. 𝕓{I} V) ∨
(0 < e ∧ ↓[d, e - 1] K ≡ L2).
-#d #e #L1 #L2 #H elim H -H d e L1 L2
+#d #e #L1 #L2 * -d e L1 L2
[ /3/
-| #L1 #L2 #I #V #e #HL12 #_ #_ #K #J #W #H destruct -L1 I V /3/
-| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #_ #H elim (plus_S_eq_O_false … H)
+| #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct -L1 I V /3/
+| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H elim (plus_S_eq_O_false … H)
]
qed.
∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 &
↑[d - 1, e] V2 ≡ V1 &
L1 = K1. 𝕓{I} V1.
-#d #e #L1 #L2 #H elim H -H d e L1 L2
+#d #e #L1 #L2 * -d e L1 L2
[ #L #H elim (lt_refl_false … H)
-| #L1 #L2 #I #V #e #_ #_ #H elim (lt_refl_false … H)
-| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #_ #I #L2 #V2 #H destruct -X Y Z;
+| #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H)
+| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct -X Y Z;
/2 width=5/
]
qed.
∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & ↑[d - 1, e] V2 ≡ V1 &
L1 = K1. 𝕓{I} V1.
/2/ qed.
+
+(* Basic properties *********************************************************)
+
+lemma drop_drop_lt: ∀L1,L2,I,V,e.
+ ↓[0, e - 1] L1 ≡ L2 → 0 < e → ↓[0, e] L1. 𝕓{I} V ≡ L2.
+#L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) /2/
+qed.
+
+lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 →
+ ↓[O, e + 1] L1 ≡ K2.
+#L1 elim L1 -L1
+[ #I2 #K2 #V2 #e #H elim (drop_inv_sort1 … H) -H #H destruct
+| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H
+ elim (drop_inv_O1 … H) -H * #He #H
+ [ -IHL1; destruct -e K2 I2 V2 /2/
+ | @drop_drop >(plus_minus_m_m e 1) /2/
+ ]
+]
+qed.
+
+lemma drop_leq_drop1: ∀L1,L2,d,e. L1 [d, e] ≈ L2 →
+ ∀I,K1,V,i. ↓[0, i] L1 ≡ K1. 𝕓{I} V →
+ d ≤ i → i < d + e →
+ ∃∃K2. K1 [0, d + e - i - 1] ≈ K2 &
+ ↓[0, i] L2 ≡ K2. 𝕓{I} V.
+#L1 #L2 #d #e #H elim H -H L1 L2 d e
+[ #d #e #I #K1 #V #i #H
+ elim (drop_inv_sort1 … H) -H #H destruct
+| #L1 #L2 #I1 #I2 #V1 #V2 #_ #_ #I #K1 #V #i #_ #_ #H
+ elim (lt_zero_false … H)
+| #L1 #L2 #I #V #e #HL12 #IHL12 #J #K1 #W #i #H #_ #Hie
+ elim (drop_inv_O1 … H) -H * #Hi #HLK1
+ [ -IHL12 Hie; destruct -i K1 J W;
+ <minus_n_O <minus_plus_m_m /2/
+ | -HL12;
+ elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 // [2: /2/ ] -Hie >arith_g1 // /3/
+ ]
+| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #I #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide
+ lapply (plus_S_le_to_pos … Hdi) #Hi
+ lapply (drop_inv_drop1 … H ?) -H // #HLK1
+ elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 [2: /2/ |3: /2/ ] -Hdi Hide >arith_g1 // /3/
+]
+qed.
(* DROPPING *****************************************************************)
-(* the main properties ******************************************************)
+(* Main properties **********************************************************)
lemma drop_conf_ge: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 →
∀e2,L2. ↓[0, e2] L ≡ L2 → d1 + e1 ≤ e2 →
]
qed.
+lemma drop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L.
+ ↓[d1, e1] L1 ≡ L → ↓[0, e2] L ≡ L2 → d1 ≤ e2 →
+ ↓[0, e2 + e1] L1 ≡ L2.
+#e1 #e1 #e2 >commutative_plus /2 width=5/
+qed.
+
axiom drop_div: ∀e1,L1,L. ↓[0, e1] L1 ≡ L → ∀e2,L2. ↓[0, e2] L2 ≡ L →
∃∃L0. ↓[0, e1] L0 ≡ L2 & ↓[e1, e2] L0 ≡ L1.
--- /dev/null
+(*
+ ||M|| This file is part of HELM, an Hypertextual, Electronic
+ ||A|| Library of Mathematics, developed at the Computer Science
+ ||T|| Department of the University of Bologna, Italy.
+ ||I||
+ ||T||
+ ||A|| This file is distributed under the terms of the
+ \ / GNU General Public License Version 2
+ \ /
+ V_______________________________________________________________ *)
+
+include "lambda-delta/syntax/lenv.ma".
+
+(* LOCAL ENVIRONMENT EQUALITY ***********************************************)
+
+inductive leq: lenv → nat → nat → lenv → Prop ≝
+| leq_sort: ∀d,e. leq (⋆) d e (⋆)
+| leq_comp: ∀L1,L2,I1,I2,V1,V2.
+ leq L1 0 0 L2 → leq (L1. 𝕓{I1} V1) 0 0 (L2. 𝕓{I2} V2)
+| leq_eq: ∀L1,L2,I,V,e. leq L1 0 e L2 → leq (L1. 𝕓{I} V) 0 (e + 1) (L2.𝕓{I} V)
+| leq_skip: ∀L1,L2,I1,I2,V1,V2,d,e.
+ leq L1 d e L2 → leq (L1. 𝕓{I1} V1) (d + 1) e (L2. 𝕓{I2} V2)
+.
+
+interpretation "local environment equality" 'Eq L1 d e L2 = (leq L1 d e L2).
+
+(* Basic properties *********************************************************)
+
+lemma leq_refl: ∀d,e,L. L [d, e] ≈ L.
+#d elim d -d
+[ #e elim e -e [ #L elim L -L /2/ | #e #IHe #L elim L -L /2/ ]
+| #d #IHd #e #L elim L -L /2/
+]
+qed.
+
+lemma leq_sym: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L2 [d, e] ≈ L1.
+#L1 #L2 #d #e #H elim H -H L1 L2 d e /2/
+qed.
+
+lemma leq_skip_lt: ∀L1,L2,d,e. leq L1 (d - 1) e L2 → 0 < d →
+ ∀I1,I2,V1,V2. L1. 𝕓{I1} V1 [d, e] ≈ L2. 𝕓{I2} V2.
+
+#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) /2/
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma leq_inv_sort1_aux: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L1 = ⋆ → L2 = ⋆.
+#L1 #L2 #d #e #H elim H -H L1 L2 d e
+[ //
+| #L1 #L2 #I1 #I2 #V1 #V2 #_ #_ #H destruct
+| #L1 #L2 #I #V #e #_ #_ #H destruct
+| #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #_ #H destruct
+qed.
+
+lemma leq_inv_sort1: ∀L2,d,e. ⋆ [d, e] ≈ L2 → L2 = ⋆.
+/2 width=5/ qed.
+
+lemma leq_inv_sort2: ∀L1,d,e. L1 [d, e] ≈ ⋆ → L1 = ⋆.
+/3/ qed.
]
qed.
+lemma lift_total: ∀T1,d,e. ∃T2. ↑[d,e] T1 ≡ T2.
+#T1 elim T1 -T1
+[ /2/
+| #i #d #e elim (lt_or_ge i d) /3/
+| * #I #V1 #T1 #IHV1 #IHT1 #d #e
+ elim (IHV1 d e) -IHV1 #V2 #HV12
+ [ elim (IHT1 (d+1) e) -IHT1 /3/
+ | elim (IHT1 d e) -IHT1 /3/
+ ]
+]
+qed.
+
(* Basic inversion lemmas ***************************************************)
lemma lift_inv_refl_aux: ∀d,e,T1,T2. ↑[d, e] T1 ≡ T2 → e = 0 → T1 = T2.
+++ /dev/null
-(**************************************************************************)
-(* ___ *)
-(* ||M|| *)
-(* ||A|| A project by Andrea Asperti *)
-(* ||T|| *)
-(* ||I|| Developers: *)
-(* ||T|| The HELM team. *)
-(* ||A|| http://helm.cs.unibo.it *)
-(* \ / *)
-(* \ / This file is distributed under the terms of the *)
-(* v GNU General Public License Version 2 *)
-(* *)
-(**************************************************************************)
-
-include "lambda-delta/substitution/lift_defs.ma".
-
-(* RELOCATION ***************************************************************)
-
-(* Functional properties ****************************************************)
-
-lemma lift_total: ∀T1,d,e. ∃T2. ↑[d,e] T1 ≡ T2.
-#T1 elim T1 -T1
-[ /2/
-| #i #d #e elim (lt_or_ge i d) /3/
-| * #I #V1 #T1 #IHV1 #IHT1 #d #e
- elim (IHV1 d e) -IHV1 #V2 #HV12
- [ elim (IHT1 (d+1) e) -IHT1 /3/
- | elim (IHT1 d e) -IHT1 /3/
- ]
-]
-qed.
-
-lemma lift_mono: ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2.
-#d #e #T #U1 #H elim H -H d e T U1
-[ #k #d #e #X #HX
- lapply (lift_inv_sort1 … HX) -HX //
-| #i #d #e #Hid #X #HX
- lapply (lift_inv_lref1_lt … HX ?) -HX //
-| #i #d #e #Hdi #X #HX
- lapply (lift_inv_lref1_ge … HX ?) -HX //
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
- elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
- elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
-]
-qed.
-
-lemma lift_inj: ∀d,e,T1,U. ↑[d,e] T1 ≡ U → ∀T2. ↑[d,e] T2 ≡ U → T1 = T2.
-#d #e #T1 #U #H elim H -H d e T1 U
-[ #k #d #e #X #HX
- lapply (lift_inv_sort2 … HX) -HX //
-| #i #d #e #Hid #X #HX
- lapply (lift_inv_lref2_lt … HX ?) -HX //
-| #i #d #e #Hdi #X #HX
- lapply (lift_inv_lref2_ge … HX ?) -HX /2/
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
- elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
-| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
- elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
-]
-qed.
(* Main properies ***********************************************************)
+lemma lift_inj: ∀d,e,T1,U. ↑[d,e] T1 ≡ U → ∀T2. ↑[d,e] T2 ≡ U → T1 = T2.
+#d #e #T1 #U #H elim H -H d e T1 U
+[ #k #d #e #X #HX
+ lapply (lift_inv_sort2 … HX) -HX //
+| #i #d #e #Hid #X #HX
+ lapply (lift_inv_lref2_lt … HX ?) -HX //
+| #i #d #e #Hdi #X #HX
+ lapply (lift_inv_lref2_ge … HX ?) -HX /2/
+| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
+| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
+]
+qed.
+
lemma lift_div_le: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T →
∀d2,e2,T2. ↑[d2 + e1, e2] T2 ≡ T →
d1 ≤ d2 →
lapply (lt_to_le_to_lt … Hid1 Hd12) -Hd12 #Hid2 /4/
| #i #d1 #e2 #Hid1 #d2 #e1 #_ #Hd21 #He12
lapply (transitive_le …(i+e1) Hd21 ?) /2/ -Hd21 #Hd21
- <(plus_plus_minus_m_m e1 e2 i) /3/
+ <(arith_d1 i e2 e1) // /3/
| #I #V1 #V2 #T1 #T2 #d1 #e2 #_ #_ #IHV #IHT #d2 #e1 #Hd12 #Hd21 #He12
elim (IHV … Hd12 Hd21 He12) -IHV #V0 #HV0a #HV0b
elim (IHT (d2+1) … ? ? He12) /3 width = 5/
]
qed.
+lemma lift_mono: ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2.
+#d #e #T #U1 #H elim H -H d e T U1
+[ #k #d #e #X #HX
+ lapply (lift_inv_sort1 … HX) -HX //
+| #i #d #e #Hid #X #HX
+ lapply (lift_inv_lref1_lt … HX ?) -HX //
+| #i #d #e #Hdi #X #HX
+ lapply (lift_inv_lref1_ge … HX ?) -HX //
+| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (lift_inv_bind1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
+| #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
+]
+qed.
+
lemma lift_trans_be: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T →
∀d2,e2,T2. ↑[d2, e2] T ≡ T2 →
d1 ≤ d2 → d2 ≤ d1 + e1 → ↑[d1, e1 + e2] T1 ≡ T2.
+++ /dev/null
-(*
- ||M|| This file is part of HELM, an Hypertextual, Electronic
- ||A|| Library of Mathematics, developed at the Computer Science
- ||T|| Department of the University of Bologna, Italy.
- ||I||
- ||T||
- ||A|| This file is distributed under the terms of the
- \ / GNU General Public License Version 2
- \ /
- V_______________________________________________________________ *)
-
-include "lambda-delta/substitution/drop_defs.ma".
-
-(* PARALLEL SUBSTITUTION ****************************************************)
-
-inductive ps: lenv → term → nat → nat → term → Prop ≝
-| ps_sort : ∀L,k,d,e. ps L (⋆k) d e (⋆k)
-| ps_lref : ∀L,i,d,e. ps L (#i) d e (#i)
-| ps_subst: ∀L,K,V,U1,U2,i,d,e.
- d ≤ i → i < d + e →
- ↓[0, i] L ≡ K. 𝕓{Abbr} V → ps K V 0 (d + e - i - 1) U1 →
- ↑[0, i + 1] U1 ≡ U2 → ps L (#i) d e U2
-| ps_bind : ∀L,I,V1,V2,T1,T2,d,e.
- ps L V1 d e V2 → ps (L. 𝕓{I} V1) T1 (d + 1) e T2 →
- ps L (𝕓{I} V1. T1) d e (𝕓{I} V2. T2)
-| ps_flat : ∀L,I,V1,V2,T1,T2,d,e.
- ps L V1 d e V2 → ps L T1 d e T2 →
- ps L (𝕗{I} V1. T1) d e (𝕗{I} V2. T2)
-.
-
-interpretation "parallel substritution" 'PSubst L T1 d e T2 = (ps L T1 d e T2).
-
-(* Basic properties *********************************************************)
-
-lemma subst_refl: ∀T,L,d,e. L ⊢ T [d, e] ≫ T.
-#T elim T -T //
-#I elim I -I /2/
-qed.
-
-(* Basic inversion lemmas ***************************************************)
-
-lemma ps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
- ∀I,V1,T1. U1 = 𝕓{I} V1. T1 →
- ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 &
- L. 𝕓{I} V1 ⊢ T1 [d + 1, e] ≫ T2 &
- U2 = 𝕓{I} V2. T2.
-#d #e #L #U1 #U2 #H elim H -H 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 #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
-]
-qed.
-
-lemma subst_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ 𝕓{I} V1. T1 [d, e] ≫ U2 →
- ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 &
- L. 𝕓{I} V1 ⊢ T1 [d + 1, e] ≫ T2 &
- U2 = 𝕓{I} V2. T2.
-/2/ qed.
-
-lemma subst_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
- ∀I,V1,T1. U1 = 𝕗{I} V1. T1 →
- ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
- U2 = 𝕗{I} V2. T2.
-#d #e #L #U1 #U2 #H elim H -H 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 #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/
-]
-qed.
-
-lemma subst_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ 𝕗{I} V1. T1 [d, e] ≫ U2 →
- ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
- U2 = 𝕗{I} V2. T2.
-/2/ qed.
+++ /dev/null
-(*
- ||M|| This file is part of HELM, an Hypertextual, Electronic
- ||A|| Library of Mathematics, developed at the Computer Science
- ||T|| Department of the University of Bologna, Italy.
- ||I||
- ||T||
- ||A|| This file is distributed under the terms of the
- \ / GNU General Public License Version 2
- \ /
- V_______________________________________________________________ *)
-
-include "lambda-delta/substitution/ps_defs.ma".
-
-(* PARALLEL SUBSTITUTION ****************************************************)
-
-axiom ps_conf: ∀L,T0,d,e,T1. L ⊢ T0 [d, e] ≫ T1 → ∀T2. L ⊢ T0 [d, e] ≫ T2 →
- ∃∃T. L ⊢ T1 [d, e] ≫ T & L ⊢ T2 [d, e] ≫ T.
--- /dev/null
+(*
+ ||M|| This file is part of HELM, an Hypertextual, Electronic
+ ||A|| Library of Mathematics, developed at the Computer Science
+ ||T|| Department of the University of Bologna, Italy.
+ ||I||
+ ||T||
+ ||A|| This file is distributed under the terms of the
+ \ / GNU General Public License Version 2
+ \ /
+ V_______________________________________________________________ *)
+
+include "lambda-delta/substitution/drop_defs.ma".
+
+(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************)
+
+inductive pts: lenv → term → nat → nat → term → Prop ≝
+| pts_sort : ∀L,k,d,e. pts L (⋆k) d e (⋆k)
+| pts_lref : ∀L,i,d,e. pts L (#i) d e (#i)
+| pts_subst: ∀L,K,V,U1,U2,i,d,e.
+ d ≤ i → i < d + e →
+ ↓[0, i] L ≡ K. 𝕓{Abbr} V → pts K V 0 (d + e - i - 1) U1 →
+ ↑[0, i + 1] U1 ≡ U2 → pts L (#i) d e U2
+| pts_bind : ∀L,I,V1,V2,T1,T2,d,e.
+ pts L V1 d e V2 → pts (L. 𝕓{I} V1) T1 (d + 1) e T2 →
+ pts L (𝕓{I} V1. T1) d e (𝕓{I} V2. T2)
+| pts_flat : ∀L,I,V1,V2,T1,T2,d,e.
+ pts L V1 d e V2 → pts L T1 d e T2 →
+ pts L (𝕗{I} V1. T1) d e (𝕗{I} V2. T2)
+.
+
+interpretation "partial telescopic substritution"
+ 'PSubst L T1 d e T2 = (pts L T1 d e T2).
+
+(* Basic properties *********************************************************)
+
+lemma pts_leq_repl: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≫ T2 →
+ ∀L2. L1 [d, e] ≈ L2 → L2 ⊢ 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
+ @pts_subst [4,5,6,8: // |1,2,3: skip | /2/ ] (**) (* /3 width=6/ is too slow *)
+| /4/
+| /3/
+]
+qed.
+
+lemma pts_refl: ∀T,L,d,e. L ⊢ T [d, e] ≫ T.
+#T elim T -T //
+#I elim I -I /2/
+qed.
+
+lemma pts_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 [d1, e1] ≫ T2 →
+ ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 →
+ L ⊢ T1 [d2, e2] ≫ 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
+ lapply (transitive_le … Hd12 … Hid1) -Hd12 Hid1 #Hid2
+ lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 #Hide2
+ @pts_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2/ ] (**) (* /4 width=6/ is too slow *)
+| /4/
+| /4/
+]
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma pts_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
+ ∀I,V1,T1. U1 = 𝕓{I} V1. T1 →
+ ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 &
+ L. 𝕓{I} V1 ⊢ T1 [d + 1, e] ≫ T2 &
+ U2 = 𝕓{I} V2. T2.
+#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 #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
+]
+qed.
+
+lemma pts_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ 𝕓{I} V1. T1 [d, e] ≫ U2 →
+ ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 &
+ L. 𝕓{I} V1 ⊢ T1 [d + 1, e] ≫ T2 &
+ U2 = 𝕓{I} V2. T2.
+/2/ qed.
+
+lemma pts_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 [d, e] ≫ U2 →
+ ∀I,V1,T1. U1 = 𝕗{I} V1. T1 →
+ ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
+ U2 = 𝕗{I} V2. T2.
+#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 #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/
+]
+qed.
+
+lemma pts_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ 𝕗{I} V1. T1 [d, e] ≫ U2 →
+ ∃∃V2,T2. L ⊢ V1 [d, e] ≫ V2 & L ⊢ T1 [d, e] ≫ T2 &
+ U2 = 𝕗{I} V2. T2.
+/2/ qed.
--- /dev/null
+(*
+ ||M|| This file is part of HELM, an Hypertextual, Electronic
+ ||A|| Library of Mathematics, developed at the Computer Science
+ ||T|| Department of the University of Bologna, Italy.
+ ||I||
+ ||T||
+ ||A|| This file is distributed under the terms of the
+ \ / GNU General Public License Version 2
+ \ /
+ V_______________________________________________________________ *)
+
+include "lambda-delta/substitution/lift_main.ma".
+include "lambda-delta/substitution/drop_main.ma".
+include "lambda-delta/substitution/pts_defs.ma".
+
+(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************)
+
+(* Relocation properties ****************************************************)
+
+lemma pts_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 →
+ dt + et ≤ d →
+ 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 //
+| #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
+ 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
+ @pts_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2 width=6/ ] (**) (* explicit constructor *)
+| #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;
+ @pts_bind [ /2 width=6/ | @IHT12 [3,4,5: /2/ |1,2: skip | /2/ ] ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *)
+| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd
+ elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
+ elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2;
+ /3 width=6/
+]
+qed.
+
+lemma pts_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 →
+ d ≤ dt →
+ 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 //
+| #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 *)
+ 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
+ @pts_subst [4,5: /2/ |6,7,8: // |1,2,3: skip ] (**) (* /3 width=8/ is too slow *)
+| #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;
+ @pts_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *)
+| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt
+ elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1
+ elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct -U1 U2;
+ /3 width=5/
+]
+qed.
+
+lemma pts_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 →
+ ∃∃T2. K ⊢ T1 [dt, et] ≫ T2 & ↑[d, e] T2 ≡ U2.
+#L #U1 #U2 #dt #et #H elim H -H L U1 U2 dt et
+[ #L #k #dt #et #K #d #e #_ #T1 #H #_
+ 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
+ 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/
+| #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 *)
+| #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 //
+ elim (IHU12 … HLK … HTU1 ?) -IHU12 HLK HTU1 // /3 width=5/
+]
+qed.
+
+lemma pts_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 →
+ ∃∃T2. K ⊢ T1 [dt - e, et] ≫ T2 & ↑[d, e] T2 ≡ U2.
+#L #U1 #U2 #dt #et #H elim H -H L U1 U2 dt et
+[ #L #k #dt #et #K #d #e #_ #T1 #H #_
+ 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
+ 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 (lift_inv_lref2_ge … H … Hdei) -H #H destruct -T1;
+ lapply (drop_conf_ge … HLK … HLKV ?) -HLK HLKV L // #HKV
+ elim (lift_free … HV12 d (i - e + 1) ? ? ?) -HV12; [2,3,4: normalize /2/ ] -Hdei >arith_e2 // #V0 #HV10 #HV02
+ @ex2_1_intro
+ [2: @pts_subst [4: /2/ |6,7,8: // |1,2,3: skip |5: @arith5 // ]
+ |1: skip
+ | //
+ ] (**) (* explicitc constructors *)
+| #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 (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
+ elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct -X;
+ elim (IHV12 … HLK … HWV1 ?) -IHV12 HWV1 //
+ elim (IHU12 … HLK … HTU1 ?) -IHU12 HLK HTU1 // /3 width=5/
+]
+qed.
+
+lemma pts_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
+[ //
+| //
+| #L #K #V #V1 #V2 #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)
+ | lapply (lt_to_le_to_lt … Hide … H) -Hide H #H
+ elim (lt_refl_false … H)
+ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #H destruct -X
+ >IHV12 // >IHT12 //
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct -X
+ >IHV12 // >IHT12 //
+]
+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) ->
+ (EX u1 | (subst0 (minus i h) v u1 u2) &
+ t1 = (lift h d u1)
+ ).
+
+
+ Theorem subst0_gen_lift_rev_lelt: (t1,v,u2:?; i,h,d:?)
+ (subst0 i v t1 (lift h d u2)) ->
+ (le d i) -> (lt i (plus d h)) ->
+ (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)).
+*)
--- /dev/null
+(*
+ ||M|| This file is part of HELM, an Hypertextual, Electronic
+ ||A|| Library of Mathematics, developed at the Computer Science
+ ||T|| Department of the University of Bologna, Italy.
+ ||I||
+ ||T||
+ ||A|| This file is distributed under the terms of the
+ \ / GNU General Public License Version 2
+ \ /
+ V_______________________________________________________________ *)
+
+include "lambda-delta/substitution/pts_split.ma".
+
+(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************)
+
+(* Main properties **********************************************************)
+
+lemma pts_trans: ∀L,T1,T,d,e. L ⊢ T1 [d, e] ≫ T → ∀T2. L ⊢ T [d, e] ≫ T2 →
+ L ⊢ T1 [d, e] ≫ T2.
+#L #T1 #T #d #e #H elim H -L T1 T d e
+[ //
+| //
+| #L #K #V #V1 #V2 #i #d #e #Hdi #Hide #HLK #_ #HV12 #IHV12 #T2 #HVT2
+ lapply (drop_fwd_drop2 … HLK) #H
+ elim (pts_inv_lift1_up … HVT2 … H … HV12 ? ? ?) -HVT2 H HV12 // normalize [2: /2/ ] #V #HV1 #HVT2
+ @pts_subst [4,5,6,8: // |1,2,3: skip | /2/ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (pts_inv_bind1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X;
+ @pts_bind /2/ @IHT12 @(pts_leq_repl … HT2) /2/
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (pts_inv_flat1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X /3/
+]
+qed.
+
+axiom pts_conf: ∀L,T0,d,e,T1. L ⊢ T0 [d, e] ≫ T1 → ∀T2. L ⊢ T0 [d, e] ≫ T2 →
+ ∃∃T. L ⊢ T1 [d, e] ≫ T & L ⊢ T2 [d, e] ≫ T.
+
+(*
+ Theorem subst0_subst0: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) ->
+ (u1,u:?; i:?) (subst0 i u u1 u2) ->
+ (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t t2)).
+
+ Theorem subst0_subst0_back: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) ->
+ (u1,u:?; i:?) (subst0 i u u2 u1) ->
+ (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t2 t)).
+
+*)
--- /dev/null
+(*
+ ||M|| This file is part of HELM, an Hypertextual, Electronic
+ ||A|| Library of Mathematics, developed at the Computer Science
+ ||T|| Department of the University of Bologna, Italy.
+ ||I||
+ ||T||
+ ||A|| This file is distributed under the terms of the
+ \ / GNU General Public License Version 2
+ \ /
+ V_______________________________________________________________ *)
+
+include "lambda-delta/substitution/pts_lift.ma".
+
+(* PARTIAL TELESCOPIC SUBSTITUTION ******************************************)
+
+(* Split properties *********************************************************)
+
+lemma pts_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → ∀i. d ≤ i → i ≤ d + e →
+ ∃∃T. L ⊢ T1 [d, i - d] ≫ T & L ⊢ T [i, d + e - i] ≫ T2.
+#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 (pts_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: @pts_lref |1: skip | /2 width=6/ ] (**) (* /3 width=6 is too slow *)
+ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+ elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2
+ elim (IHT12 (i + 1) ? ?) -IHT12 [2: /2 by arith4/ |3: /2/ ] (* just /2/ is too slow *)
+ -Hdi Hide >arith_c1 >arith_c1x #T #HT1 #HT2
+ @ex2_1_intro [2,3: @pts_bind | skip ] (**) (* explicit constructors *)
+ [3: @HV1 |4: @HT1 |5: // |1,2: skip | /3 width=5/ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide
+ elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 //
+ -Hdi Hide /3 width=5/
+]
+qed.
+
+lemma pts_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 [dt, et] ≫ U2 →
+ ∀K,d,e. ↓[d, e] L ≡ K → ∀T1. ↑[d, e] T1 ≡ U1 →
+ d ≤ dt → dt ≤ d + e → d + e ≤ dt + et →
+ ∃∃T2. K ⊢ T1 [d, dt + et - (d + e)] ≫ T2 & ↑[d, e] T2 ≡ U2.
+#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet
+elim (pts_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2
+lapply (pts_weak … HU1 d e ? ?) -HU1 // <plus_minus_m_m_comm // -Hddt Hdtde #HU1
+lapply (pts_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -U1;
+elim (pts_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 HLK HTU1 // <minus_plus_m_m /2/
+qed.
<key name="ex">7 6</key>
<key name="or">3</key>
<key name="or">4</key>
+ <key name="and">3</key>
</section>
</helm_registry>
interpretation "multiple disjunction connective (4)" 'Or P0 P1 P2 P3 = (or4 P0 P1 P2 P3).
+(* multiple conjunction connective (3) *)
+
+inductive and3 (P0,P1,P2:Prop) : Prop ≝
+ | and3_intro: P0 → P1 → P2 → and3 ? ? ?
+.
+
+interpretation "multiple conjunction connective (3)" 'And P0 P1 P2 = (and3 P0 P1 P2).
+
non associative with precedence 30
for @{ 'Or $P0 $P1 $P2 $P3 }.
+(* multiple conjunction connective (3) *)
+
+notation "hvbox(∧∧ term 34 P0 break & term 34 P1 break & term 34 P2)"
+ non associative with precedence 35
+ for @{ 'And $P0 $P1 $P2 }.
+