∀T. ⦃G, L1⦄ ⊢ ⬊*[h, g] T → ⦃G, L2⦄ ⊢ ⬊*[h, g] T.
#h #g #G #L1 #L2 #HL12 #T #H @(csx_ind_alt … H) -T
/4 width=3 by csx_intro, lpx_cpx_trans/
-qed.
+qed-.
lemma csx_abst: ∀h,g,a,G,L,W. ⦃G, L⦄ ⊢ ⬊*[h, g] W →
∀T. ⦃G, L.ⓛW⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓛ{a}W.T.
--- /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 "basic_2/computation/csx_lpx.ma".
+include "basic_2/computation/lpxs.ma".
+
+(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
+
+(* Properties on sn extended parallel computation for local environments ****)
+
+lemma csx_lpxs_conf: ∀h,g,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
+ ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, g] T → ⦃G, L2⦄ ⊢ ⬊*[h, g] T.
+#h #g #G #L1 #L2 #H @(lpxs_ind … H) -L2 /3 by lpxs_strap1, csx_lpx_conf/
+qed-.
include "basic_2/computation/lsx_csx.ma".
include "basic_2/computation/fsb_alt.ma".
-axiom lsx_fqup_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃+ ⦃G2, L2, T2⦄ →
- G1 ⊢ ⋕⬊*[h, g, T1, 0] L1 → G2 ⊢ ⋕⬊*[h, g, T2, 0] L2.
-
axiom fqup_lpxs_trans_nlleq: ∀h,g,G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊃+ ⦃G2, K2, T2⦄ →
∀L2. ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 → (K2 ⋕[T2, 0] L2 →⊥) →
∃∃L1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 &
lemma csx_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⦥[h, g] T.
#h #g #G1 #L1 #T1 #H @(csx_ind_alt … H) -T1
-#T1 #HT1 @(lsx_ind h g G1 T1 0 … L1) /2 width=1 by csx_lsx/ -L1
+#T1 #HT1 @(csx_ind_lsx … (yinj 0) … HT1) -L1
#L1 @(fqup_wf_ind … G1 L1 T1) -G1 -L1 -T1
#G1 #L1 #T1 #IHu #H1 #IHl #IHc @fsb_intro
#G2 #L2 #T2 * -G2 -L2 -T2
-[ #G2 #L2 #T2 #H12 @IHu -IHu /2 width=5 by lsx_fqup_conf/ -H1 [| -IHl ]
+[ #G2 #L2 #T2 #H12 @IHu -IHu // /2 width=5 by csx_fqup_conf/ -H1 [| -IHl ]
[ #L0 #HL20 #HnL20 #_ elim (fqup_lpxs_trans_nlleq … H12 … HL20 HnL20) -L2
/6 width=5 by fsb_fpbs_trans, lpxs_fpbs, fqup_fpbs, lpxs_cpxs_trans/
| #T0 #HT20 #HnT20 elim (fqup_cpxs_trans_neq … H12 … HT20 HnT20) -T2
(* *)
(**************************************************************************)
-include "basic_2/computation/csx_alt.ma".
-include "basic_2/computation/csx_lift.ma".
+include "basic_2/computation/csx_lpxs.ma".
include "basic_2/computation/lcosx_cpxs.ma".
(* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
]
qed.
+
+(* Advanced eliminators *****************************************************)
+
+fact csx_ind_lsx_aux: ∀h,g,G,T,d. ∀R:predicate lenv.
+ (∀L1. ⦃G, L1⦄ ⊢ ⬊*[h, g] T →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ⋕[T, d] L2 → ⊥) → R L2) →
+ R L1
+ ) →
+ ∀L. G ⊢ ⋕⬊*[h, g, T, d] L → ⦃G, L⦄ ⊢ ⬊*[h, g] T → R L.
+#h #g #G #T #d #R #IH #L #H @(lsx_ind … H) -L
+#L1 #_ #IHL1 #HL1 @IH -IH //
+#L2 #HL12 #HnT @IHL1 -IHL1 /2 width=3 by csx_lpxs_conf/
+qed-.
+
+lemma csx_ind_lsx: ∀h,g,G,T,d. ∀R:predicate lenv.
+ (∀L1. ⦃G, L1⦄ ⊢ ⬊*[h, g] T →
+ (∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → (L1 ⋕[T, d] L2 → ⊥) → R L2) →
+ R L1
+ ) →
+ ∀L. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R L.
+#h #g #G #T #d #R #IH #L #HL @(csx_ind_lsx_aux … d … HL) /4 width=1 by csx_lsx/
+qed-.
[ "lcosx ( ? ⊢ ⧤⬊*[?,?,?] ? )" "lcosx_cpxs" * ]
[ "lsx ( ? ⊢ ⋕⬊*[?,?,?,?] ? )" "lsx_ldrop" + "lsx_cpxs" + "lsx_csx" * ]
[ "csx_vector ( ⦃?,?⦄ ⊢ ⬊*[?,?] ? )" "csx_tstc_vector" + "csx_aaa" * ]
- [ "csx ( ⦃?,?⦄ ⊢ ⬊*[?,?] ? )" "csx_alt ( ⦃?,?⦄ ⊢ ⬊⬊*[?,?] ? )" "csx_lift" + "csx_lpx" + "csx_fpbs" * ]
+ [ "csx ( ⦃?,?⦄ ⊢ ⬊*[?,?] ? )" "csx_alt ( ⦃?,?⦄ ⊢ ⬊⬊*[?,?] ? )" "csx_lift" + "csx_lpx" + "csx_lpxs" + "csx_fpbs" * ]
}
]
[ { "\"big tree\" parallel computation" * } {
--- /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 *)
+(* *)
+(**************************************************************************)
+
+(* GENERAL NOTATION USED BY THE FORMAL SYSTEM λδ ****************************)
+
+notation "hvbox( x ⊞ break term 46 y1 ⊟ break term 46 y2 ≡ break term 46 z )"
+ non associative with precedence 45
+ for @{ 'RPlusMinus $x $y1 $y2 $z }.
--- /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 "ground_2/notation/relations/rplusminus_4.ma".
+include "ground_2/ynat/ynat_plus.ma".
+
+(* NATURAL NUMBERS WITH INFINITY ********************************************)
+
+(* algebraic x + y1 - y2 = z *)
+inductive yrpm (x:ynat) (y1:ynat) (y2:ynat): predicate ynat ≝
+| yrpm_ge: y2 ≤ y1 → yrpm x y1 y2 (x + (y1 - y2))
+| yrpm_lt: y1 < y2 → yrpm x y1 y2 (x - (y2 - y1))
+.
+
+interpretation "ynat 'algebraic plus-minus' (relational)"
+ 'RPlusMinus x y1 y2 z = (yrpm x y1 y2 z).
+
+(* Basic inversion lemmas ***************************************************)
+
+lemma ypm_inv_ge: ∀x,y1,y2,z. x ⊞ y1 ⊟ y2 ≡ z →
+ y2 ≤ y1 → z = x + (y1 - y2).
+#x #y1 #y2 #z * -z //
+#Hy12 #H elim (ylt_yle_false … H) -H //
+qed-.
+
+lemma ypm_inv_lt: ∀x,y1,y2,z. x ⊞ y1 ⊟ y2 ≡ z →
+ y1 < y2 → z = x - (y2 - y1).
+#x #y1 #y2 #z * -z //
+#Hy21 #H elim (ylt_yle_false … H) -H //
+qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma ypm_inv_le: ∀x,y1,y2,z. x ⊞ y1 ⊟ y2 ≡ z →
+ y1 ≤ y2 → z = x - (y2 - y1).
+#x #y1 #y2 #z #H #Hy12 elim (yle_split_eq … Hy12) -Hy12 #Hy12
+[ /2 width=1 by ypm_inv_lt/
+| >(ypm_inv_ge … H) -H // destruct >yminus_refl //
+]
+qed-.
]
class "yellow"
[ { "natural numbers with infinity" * } {
- [ "ynat ( ∞ )" "ynat_pred ( ⫰? )" "ynat_succ ( ⫯? )" "ynat_le ( ?≤? )" "ynat_lt ( ?<? )" "ynat_minus ( ? - ? )" "ynat_plus ( ? + ? )" "ynat_max" "ynat_min" * ]
+ [ "ynat ( ∞ )" "ynat_pred ( ⫰? )" "ynat_succ ( ⫯? )"
+ "ynat_le ( ?≤? )" "ynat_lt ( ?<? )"
+ "ynat_minus ( ? - ? )" "ynat_plus ( ? + ? )"
+ "ynat_max" "ynat_min" * ]
}
]
class "orange"
[":"; "⁝"; ];
["."; "•"; "◦"; ];
["#"; "♯"; "⋕"; "⧤"; "⌘"; ];
- ["-"; "÷"; "⊢"; "⊩"; ];
+ ["+"; "⊞"; ];
+ ["-"; "÷"; "⊢"; "⊩"; "⊟"; ];
["="; "≝"; "⊜"; "≡"; "≃"; "≈"; "≅"; "≐"; "≑"; "≚"; "≙"; "⌆"; ];
["→"; "↦"; "⇝"; "⤞"; "⇾"; "⤍"; "⤏"; "⤳"; ] ;
["⇒"; "⤇"; "➾"; "⇨"; "➡"; "➤"; "➸"; "⇉"; "⥰"; ] ;