]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/etc/ynat_old/ynat_pred.etc
milestone update in ground, partial commit
[helm.git] / matita / matita / contribs / lambdadelta / ground / etc / ynat_old / ynat_pred.etc
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ground_2/ynat/ynat.ma".
16
17 (* INFINITARY NATURAL NUMBERS ***********************************************)
18
19 (* the predecessor on ynat *)
20 definition ypred: ynat → ynat ≝ λm. match m with
21 [ YO   ⇒ 0
22 | YS n ⇒ n
23 ].
24
25 notation "hvbox( ⫰ term 55 T )" 
26    non associative with precedence 55
27    for @{ 'Predecessor $T }.
28
29 interpretation "ynat predecessor" 'Predecessor m = (ypred m).
30
31 (* Properties ***************************************************************)
32
33 lemma ypred_S: ∀m. ⫰⫯m = m.
34 // qed.
35
36 (* Inversion lemmas *********************************************************)
37
38 lemma YS_inj: ∀m,n. ⫯m = ⫯n → m = n.
39 #m #n #H <(ypred_S m) <(ypred_S n) //
40 qed-.