]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_succ.ma
- natural numbers with infinity for lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / ynat / ynat_succ.ma
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/notation/functions/successor_1.ma".
16 include "ground_2/ynat/ynat_pred.ma".
17
18 (* NATURAL NUMBERS WITH INFINITY ********************************************)
19
20 (* the successor function *)
21 definition ysucc: ynat → ynat ≝ λm. match m with
22 [ yinj m ⇒ S m
23 | Y      ⇒ Y
24 ].
25
26 interpretation "ynat successor" 'Successor m = (ysucc m).
27
28 (* Properties ***************************************************************)
29
30 lemma ypred_succ: ∀m. ⫰⫯m = m.
31 * // qed.
32
33 (* Inversion lemmas *********************************************************)
34
35 lemma ysucc_inj: ∀m,n. ⫯m = ⫯n → m = n.
36 #m #n #H <(ypred_succ m) <(ypred_succ n) //
37 qed-.
38
39 lemma ysucc_inv_refl: ∀m. ⫯m = m → m = ∞.
40 * // normalize
41 #m #H lapply (yinj_inj … H) -H (**) (* destruct lemma needed *)
42 #H elim (lt_refl_false m) //
43 qed-.
44
45 lemma ysucc_inv_inj_sn: ∀m2,n1. yinj m2 = ⫯n1 →
46                         ∃∃m1. n1 = yinj m1 & m2 = S m1.
47 #m2 * normalize
48 [ #n1 #H destruct /2 width=3 by ex2_intro/
49 | #H destruct
50 ]
51 qed-.
52
53 lemma ysucc_inv_inj_dx: ∀m2,n1. ⫯n1 = yinj m2  →
54                         ∃∃m1. n1 = yinj m1 & m2 = S m1.
55 /2 width=1 by ysucc_inv_inj_sn/ qed-.