XOA = xoa.native
CONF = xoa.conf.xml
-TARGETS = xoa_natation.ma xoa_defs.ma
+TARGETS = xoa_natation.ma xoa.ma
all: $(TARGETS)
>(commutative_plus p) <plus_minus_m_m //
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
-[ #p #n #H >(le_O_to_eq_O … H) -H //
-| #m #IHm #p elim p -p //
- #p #_ #n #Hpm <plus_n_Sm @IHm /2/
+lemma minus_le_minus_minus_comm: ∀b,c,a. c ≤ b → a - (b - c) = a + c - b.
+#b elim b -b
+[ #c #a #H >(le_O_to_eq_O … H) -H //
+| #b #IHb #c elim c -c //
+ #c #_ #a #Hcb
+ lapply (le_S_S_to_le … Hcb) -Hcb #Hcb
+ <plus_n_Sm normalize /2/
]
qed.
\ /
V_______________________________________________________________ *)
-include "lambda-delta/syntax/length.ma".
include "lambda-delta/reduction/tpr.ma".
(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2.
/2/ qed.
-axiom cpr_pts: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2.
+lemma cpr_tps: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2.
+/3 width=5/ qed.
lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T.
/2/ qed.
#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
qed.
-axiom cpr_delta: ∀L,K,V0,V,i.
- ↓[0, i] L ≡ K. 𝕓{Abbr} V0 → ↑[0, i + 1] V0 ≡ V → L ⊢ #i ⇒ V.
-(*
-#L #K #V0 #V #i #HLK #HV0
-@ex2_1_intro [2: // |1: skip ]
-@pts_subst [4,6,7,8: // |1,2,3: skip |5:
-*)
+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 *)
+qed.
lemma cpr_cast: ∀L,V,T1,T2.
L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{Cast} V. T1 ⇒ T2.
\ /
V_______________________________________________________________ *)
-include "lambda-delta/substitution/pts.ma".
+include "lambda-delta/substitution/tps.ma".
(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
(* *)
(**************************************************************************)
-include "lambda-delta/substitution/pts_lift.ma".
+include "lambda-delta/substitution/tps_lift.ma".
include "lambda-delta/reduction/tpr.ma".
(* Relocation properties ****************************************************)
elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct -X2;
elim (lift_total T2 (d + 1) e) #U2 #HTU2
@tpr_delta
- [4: @(pts_lift_le … HT2 … HTU2 HTU0 ?) /2/ |1: skip |2: /2/ |3: /2/ ] (**) (*/3. is too slow *)
+ [4: @(tps_lift_le … HT2 … HTU2 HTU0 ?) /2/ |1: skip |2: /2/ |3: /2/ ] (**) (*/3. is too slow *)
| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct -X1;
elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct -X;
elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct -X;
elim (IHV12 … HWV1) -IHV12 HWV1 #W2 #HWV2 #HW12
elim (IHT12 … HUT1) -IHT12 HUT1 #U2 #HUT2 #HU12
- elim (pts_inv_lift1_le … HT20 … HUT2 ?) -HT20 HUT2 // [3: /2 width=5/ |2: skip ] #U0 #HU20 #HUT0
+ elim (tps_inv_lift1_le … HT20 … HUT2 ?) -HT20 HUT2 // [3: /2 width=5/ |2: skip ] #U0 #HU20 #HUT0
@ex2_1_intro [2: /2/ |1: skip |3: /2/ ] (**) (* /3 width=5/ is slow *)
| #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #X #HX
elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X;
+++ /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.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_weight.ma".
-include "lambda-delta/substitution/pts_pts.ma".
+include "lambda-delta/substitution/tps_tps.ma".
include "lambda-delta/reduction/tpr_lift.ma".
-include "lambda-delta/reduction/tpr_pts.ma".
+include "lambda-delta/reduction/tpr_tps.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_pts_bind … HV2 HT20 … HT2) -HT20 HT2 /3 width=5/
+elim (tpr_tps_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_pts_bind … HWW2 HUUU2 … HUU2) -HUU2 HUUU2 #UUU #HUUU2 #HUUU1
+ elim (tpr_tps_bind … HWW2 HUUU2 … HUU2) -HUU2 HUUU2 #UUU #HUUU2 #HUUU1
@ex2_1_intro
[2: @tpr_theta [6: @HVV |7: @HWW2 |8: @HUUU2 |1,2,3,4: skip | // ]
|1:skip
#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_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
+elim (tpr_tps_bind … HV1 HT1 … HTT1) -HT1 HTT1 #U1 #TTU1 #HTU1
+elim (tpr_tps_bind … HV2 HT2 … HTT2) -HT2 HTT2 #U2 #TTU2 #HTU2
+elim (tps_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 (pts_inv_lift1_eq … HTT1 … HTT21) -HTT1 #HTT1 destruct -T1;
+lapply (tps_inv_lift1_eq … HTT1 … HTT21) -HTT1 #HTT1 destruct -T1;
lapply (tw_lift … HTT20) -HTT20 #HTT20
elim (IH … HTX2 … HTT2) -HTX2 HTT2 IH /3/
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.ma".
+
+(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
+
+axiom tpr_tps_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_tps_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.
#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 lapply (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 →
elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 [2: /2/ |3: /2/ ] -Hdi Hide >arith_g1 // /3/
]
qed.
+
+(* Basic forvard lemmas *****************************************************)
+
+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 lapply (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_fwd_drop2_length: ∀L1,I2,K2,V2,e.
+ ↓[0, e] L1 ≡ K2. 𝕓{I2} V2 → e < |L1|.
+#L1 elim L1 -L1
+[ #I2 #K2 #V2 #e #H lapply (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 //
+ | lapply (IHL1 … H) -IHL1 H #HeK1 whd in ⊢ (? ? %) /2/
+ ]
+]
+qed.
+
+lemma drop_fwd_O1_length: ∀L1,L2,e. ↓[0, e] L1 ≡ L2 → |L2| = |L1| - e.
+#L1 elim L1 -L1
+[ #L2 #e #H >(drop_inv_sort1 … H) -H //
+| #K1 #I1 #V1 #IHL1 #L2 #e #H
+ elim (drop_inv_O1 … H) -H * #He #H
+ [ -IHL1; destruct -e L2 //
+ | lapply (IHL1 … H) -IHL1 H #H >H -H; normalize
+ >minus_le_minus_minus_comm //
+ ]
+]
+qed.
[ -He2d IHL12 #H1 #H2 destruct -e2 L /3 width=5/
| -HL12 HV12 #He2 #HL2
elim (IHL12 … HL2 ?) -IHL12 HL2 L2
- [ <minus_le_minus_minus_comm /3/ | /2/ ]
+ [ >minus_le_minus_minus_comm // /3/ | /2/ ]
]
]
qed.
\ /
V_______________________________________________________________ *)
-include "lambda-delta/syntax/lenv.ma".
+include "lambda-delta/syntax/length.ma".
(* LOCAL ENVIRONMENT EQUALITY ***********************************************)
#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 →
+lemma leq_skip_lt: ∀L1,L2,d,e. 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.
+lemma leq_fwd_length: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → |L1| = |L2|.
+#L1 #L2 #d #e #H elim H -H L1 L2 d e; normalize //
+qed.
+
(* Basic inversion lemmas ***************************************************)
lemma leq_inv_sort1_aux: ∀L1,L2,d,e. L1 [d, e] ≈ L2 → L1 = ⋆ → L2 = ⋆.
+++ /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.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_lift.ma".
-include "lambda-delta/substitution/drop_drop.ma".
-include "lambda-delta/substitution/pts.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_split … 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.
--- /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.ma".
+
+(* PARTIAL 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_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)
+| tps_flat : ∀L,I,V1,V2,T1,T2,d,e.
+ tps L V1 d e V2 → tps L T1 d e T2 →
+ tps L (𝕗{I} V1. T1) d e (𝕗{I} V2. T2)
+.
+
+interpretation "partial telescopic substritution"
+ 'PSubst L T1 d e T2 = (tps L T1 d e T2).
+
+(* Basic properties *********************************************************)
+
+lemma tps_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
+ @tps_subst [4,5,6,8: // |1,2,3: skip | /2/ ] (**) (* /3 width=6/ is too slow *)
+| /4/
+| /3/
+]
+qed.
+
+lemma tps_refl: ∀T,L,d,e. L ⊢ T [d, e] ≫ T.
+#T elim T -T //
+#I elim I -I /2/
+qed.
+
+lemma tps_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
+ @tps_subst [4,5,6,8: // |1,2,3: skip | @IHV12 /2/ ] (**) (* /4 width=6/ is too slow *)
+| /4/
+| /4/
+]
+qed.
+
+lemma tps_weak_top: ∀L,T1,T2,d,e.
+ L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 [d, |L| - d] ≫ T2.
+#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
+ 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/
+| normalize /2/
+| /2/
+]
+qed.
+
+lemma tps_weak_all: ∀L,T1,T2,d,e.
+ L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 [0, |L|] ≫ T2.
+#L #T1 #T #d #e #HT12
+lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12
+lapply (tps_weak_top … HT12) //
+qed.
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma tps_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 tps_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 tps_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 tps_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_lift.ma".
+include "lambda-delta/substitution/drop_drop.ma".
+include "lambda-delta/substitution/tps.ma".
+
+(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
+
+(* Relocation properties ****************************************************)
+
+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 →
+ 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
+ @tps_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;
+ @tps_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 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 →
+ 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
+ @tps_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;
+ @tps_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 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 →
+ ∃∃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 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 →
+ ∃∃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_split … HV12 d (i - e + 1) ? ? ?) -HV12; [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 // ]
+ |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 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
+[ //
+| //
+| #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/tps_lift.ma".
+
+(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
+
+(* Split properties *********************************************************)
+
+lemma tps_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 (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 #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: @tps_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 tps_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 (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2
+lapply (tps_weak … HU1 d e ? ?) -HU1 // <plus_minus_m_m_comm // -Hddt Hdtde #HU1
+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.
--- /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/tps_split.ma".
+
+(* PARTIAL SUBSTITUTION ON TERMS ********************************************)
+
+(* Main properties **********************************************************)
+
+lemma tps_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 (tps_inv_lift1_up … HVT2 … H … HV12 ? ? ?) -HVT2 H HV12 // normalize [2: /2/ ] #V #HV1 #HVT2
+ @tps_subst [4,5,6,8: // |1,2,3: skip | /2/ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (tps_inv_bind1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X;
+ @tps_bind /2/ @IHT12 @(tps_leq_repl … HT2) /2/
+| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
+ elim (tps_inv_flat1 … HX) -HX #V #T #HV2 #HT2 #HX destruct -X /3/
+]
+qed.
+
+axiom tps_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)).
+
+*)
<helm_registry>
<section name="matita">
<key name="rt_base_dir">$(MATITA_RT_BASE_DIR)</key>
-<!--
+<!--
<key name="system">false</key>
<key name="map_unicode_to_tex">false</key>
<key name="do_heavy_checks">true</key>
<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 }.
-