]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2A/ynat/ynat_succ.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / ground_2A / 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_2A/notation/functions/successor_1.ma".
16 include "ground_2A/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 lemma ynat_cases: ∀n:ynat. n = 0 ∨ ∃m. n = ⫯m.
34 *
35 [ * /2 width=1 by or_introl/
36   #n @or_intror @(ex_intro … n) // (**) (* explicit constructor *)
37 | @or_intror @(ex_intro … (∞)) // (**) (* explicit constructor *)
38 ]
39 qed-.
40
41 lemma ysucc_iter_Y: ∀m. ysucc^m (∞) = ∞.
42 #m elim m -m //
43 #m #IHm whd in ⊢ (??%?); >IHm //
44 qed.
45
46 (* Inversion lemmas *********************************************************)
47
48 lemma ysucc_inj: ∀m,n. ⫯m = ⫯n → m = n.
49 #m #n #H <(ypred_succ m) <(ypred_succ n) //
50 qed-.
51
52 lemma ysucc_inv_refl: ∀m. ⫯m = m → m = ∞.
53 * // normalize
54 #m #H lapply (yinj_inj … H) -H (**) (* destruct lemma needed *)
55 #H elim (lt_refl_false m) //
56 qed-.
57
58 lemma ysucc_inv_inj_sn: ∀m2,n1. yinj m2 = ⫯n1 →
59                         ∃∃m1. n1 = yinj m1 & m2 = S m1.
60 #m2 * normalize
61 [ #n1 #H destruct /2 width=3 by ex2_intro/
62 | #H destruct
63 ]
64 qed-.
65
66 lemma ysucc_inv_inj_dx: ∀m2,n1. ⫯n1 = yinj m2  →
67                         ∃∃m1. n1 = yinj m1 & m2 = S m1.
68 /2 width=1 by ysucc_inv_inj_sn/ qed-.
69
70 lemma ysucc_inv_Y_sn: ∀m. ∞ = ⫯m → m = ∞.
71 * // normalize
72 #m #H destruct
73 qed-.
74
75 lemma ysucc_inv_Y_dx: ∀m. ⫯m = ∞ → m = ∞.
76 /2 width=1 by ysucc_inv_Y_sn/ qed-.
77
78 lemma ysucc_inv_O_sn: ∀m. yinj 0 = ⫯m → ⊥. (**) (* explicit coercion *)
79 #m #H elim (ysucc_inv_inj_sn … H) -H
80 #n #_ #H destruct
81 qed-.
82
83 lemma ysucc_inv_O_dx: ∀m. ⫯m = 0 → ⊥.
84 /2 width=2 by ysucc_inv_O_sn/ qed-.