--- /dev/null
+
+lemma zero_or_gt: ∀n. n = 0 ∨ 0 < n.
+#n elim (lt_or_eq_or_gt 0 n) /2/
+#H elim (lt_zero_false … H)
+qed.
+
+lemma pippo: ∀x,y,z. x < z → y < z - x → x + y < z.
+/3 width=2/
+
+lemma arith_b1x: ∀e,x1,x2,y. y ≤ x2 → x2 ≤ x1 →
+ e + (x1 - y) - (x2 - y) = e + (x1 - x2).
+#e #x1 #x2 #y #H1 #H2
+<(arith_b1 … H1) >le_plus_minus // /2 width=1/
+qed-.
+
+lemma arith1: ∀x,y,z,w. z < y → x + (y + w) - z = x + y - z + w.
+#x #y #z #w #H <le_plus_minus_comm // /3 width=1 by lt_to_le, le_plus_a/
+qed-.
--- /dev/null
+let rec append A (l1: list A) l2 on l1 ≝
+ match l1 with
+ [ nil ⇒ l2
+ | cons hd tl ⇒ hd :: append A tl l2
+ ].
+
+interpretation "append (list)" 'Append l1 l2 = (append ? l1 l2).
--- /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 "arithmetics/nat.ma".
+
+(* INFINITARY NATURAL NUMBERS ***********************************************)
+
+(* the type of infinitary natural numbers *)
+coinductive ynat: Type[0] ≝
+| YO: ynat
+| YS: ynat → ynat
+.
+
+interpretation "ynat successor" 'Successor m = (YS m).
+
+(* the coercion of nat to ynat *)
+let rec ynat_of_nat n ≝ match n with
+[ O ⇒ YO
+| S m ⇒ YS (ynat_of_nat m)
+].
+
+coercion ynat_of_nat.
+
+(* the infinity *)
+let corec Y : ynat ≝ ⫯Y.
+
+interpretation "ynat infinity" 'Infinity = Y.
+
+(* destructing identity on ynat *)
+definition yid: ynat → ynat ≝ λm. match m with
+[ YO ⇒ 0
+| YS n ⇒ ⫯n
+].
+
+(* Properties ***************************************************************)
+
+fact yid_rew: ∀n. yid n = n.
+* // qed-.
+
+lemma Y_rew: ⫯∞ = ∞.
+<(yid_rew ∞) in ⊢ (???%); //
+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 "ground_2/notation.ma".
+include "ground_2/xoa_props.ma".
+include "ground_2/ynat/ynat.ma".
+
+(* INFINITARY NATURAL NUMBERS ***********************************************)
+
+(* "is_zero" predicate *)
+definition yzero: predicate ynat ≝ λx. match x with
+[ YO ⇒ ⊤
+| YS _ ⇒ ⊥
+].
+
+(* Inversion lemmas *********************************************************)
+
+lemma discr_YS_YO: ∀n. ⫯n = 0 → ⊥.
+#n #H change with (yzero (⫯n))
+>H -H //
+qed-.
+
+lemma discr_YO_YS: ∀n. ynat_of_nat 0 = ⫯n → ⊥. (**) (* explicit coercion *)
+/2 width=2 by discr_YS_YO/ 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 "ground_2/star.ma".
+include "ground_2/ynat/ynat_iszero.ma".
+include "ground_2/ynat/ynat_pred.ma".
+
+(* INFINITARY NATURAL NUMBERS ***********************************************)
+
+(* order relation *)
+coinductive yle: relation ynat ≝
+| yle_O: ∀n. yle 0 n
+| yle_S: ∀m,n. yle m n → yle (⫯m) (⫯n)
+.
+
+interpretation "natural 'less or equal to'" 'leq x y = (yle x y).
+
+(* Inversion lemmas *********************************************************)
+
+fact yle_inv_O2_aux: ∀m,x. m ≤ x → x = 0 → m = 0.
+#m #x * -m -x //
+#m #x #_ #H elim (discr_YS_YO … H) (**) (* destructing lemma needed *)
+qed-.
+
+lemma yle_inv_O2: ∀m. m ≤ 0 → m = 0.
+/2 width =3 by yle_inv_O2_aux/ qed-.
+
+fact yle_inv_S1_aux: ∀x,y. x ≤ y → ∀m. x = ⫯m → ∃∃n. m ≤ n & y = ⫯n.
+#x #y * -x -y
+[ #y #m #H elim (discr_YO_YS … H) (**) (* destructing lemma needed *)
+| #x #y #Hxy #m #H destruct /2 width=3 by ex2_intro/
+]
+qed-.
+
+lemma yle_inv_S1: ∀m,y. ⫯m ≤ y → ∃∃n. m ≤ n & y = ⫯n.
+/2 width=3 by yle_inv_S1_aux/ qed-.
+
+lemma yle_inv_S: ∀m,n. ⫯m ≤ ⫯n → m ≤ n.
+#m #n #H elim (yle_inv_S1 … H) -H
+#x #Hx #H destruct //
+qed-.
+
+(* Properties ***************************************************************)
+
+let corec yle_refl: reflexive … yle ≝ ?.
+* [ @yle_O | #x @yle_S // ]
+qed.
+
+let corec yle_Y: ∀m. m ≤ ∞ ≝ ?.
+* [ @yle_O | #m <Y_rew @yle_S // ]
+qed.
+
+let corec yle_S_dx: ∀m,n. m ≤ n → m ≤ ⫯n ≝ ?.
+#m #n * -m -n [ #n @yle_O | #m #n #H @yle_S /2 width=1 by/ ]
+qed.
+
+lemma yle_refl_S_dx: ∀x. x ≤ ⫯x.
+/2 width=1 by yle_refl, yle_S_dx/ qed.
+
+lemma yle_pred_sn: ∀m,n. m ≤ n → ⫰m ≤ n ≝ ?.
+* // #m #n #H elim (yle_inv_S1 … H) -H
+#x #Hm #H destruct /2 width=1 by yle_S_dx/
+qed.
+
+lemma yle_refl_pred_sn: ∀x. ⫰x ≤ x.
+/2 width=1 by yle_refl, yle_pred_sn/ qed.
+
+let corec yle_trans: Transitive … yle ≝ ?.
+#x #y * -x -y [ #x #z #_ @yle_O ]
+#x #y #Hxy #z #H elim (yle_inv_S1 … H) -H
+#n #Hyz #H destruct
+@yle_S @(yle_trans … Hxy … Hyz) (**) (* cofix not guarded by constructors *)
+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 "ground_2/ynat/ynat.ma".
+
+(* INFINITARY NATURAL NUMBERS ***********************************************)
+
+(* the predecessor on ynat *)
+definition ypred: ynat → ynat ≝ λm. match m with
+[ YO ⇒ 0
+| YS n ⇒ n
+].
+
+notation "hvbox( ⫰ term 55 T )"
+ non associative with precedence 55
+ for @{ 'Predecessor $T }.
+
+interpretation "ynat predecessor" 'Predecessor m = (ypred m).
+
+(* Properties ***************************************************************)
+
+lemma ypred_S: ∀m. ⫰⫯m = m.
+// qed.
+
+(* Inversion lemmas *********************************************************)
+
+lemma YS_inj: ∀m,n. ⫯m = ⫯n → m = n.
+#m #n #H <(ypred_S m) <(ypred_S n) //
+qed-.
["?"; "¿"; "⸮"; ];
[":"; "⁝"; ];
["."; "•"; "◦"; ];
- ["#"; "♯"; "⋕"; "⧤"; "⌘"; ];
+ ["#"; "â\99¯"; "â\8b\95"; "â§£"; "⧤"; "â\8c\98"; ];
["+"; "⊞"; ];
["-"; "÷"; "⊢"; "⊩"; "⊟"; ];
["="; "≝"; "⊜"; "≡"; "≃"; "≈"; "≅"; "≐"; "≑"; "≚"; "≙"; "⌆"; ];