1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground_2/ynat/ynat.ma".
17 (* INFINITARY NATURAL NUMBERS ***********************************************)
19 (* the predecessor on ynat *)
20 definition ypred: ynat → ynat ≝ λm. match m with
25 notation "hvbox( ⫰ term 55 T )"
26 non associative with precedence 55
27 for @{ 'Predecessor $T }.
29 interpretation "ynat predecessor" 'Predecessor m = (ypred m).
31 (* Properties ***************************************************************)
33 lemma ypred_S: ∀m. ⫰⫯m = m.
36 (* Inversion lemmas *********************************************************)
38 lemma YS_inj: ∀m,n. ⫯m = ⫯n → m = n.
39 #m #n #H <(ypred_S m) <(ypred_S n) //