]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_succ.ma
- subtraction (and related notions) removed
[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 lemma ysucc_inj: ∀m:nat. ⫯m = S m.
29 // qed.
30
31 lemma ysucc_Y: ⫯(∞) = ∞.
32 // qed.
33
34 (* Properties ***************************************************************)
35
36 lemma ypred_succ: ∀m. ⫰⫯m = m.
37 * // qed.
38
39 lemma ynat_cases: ∀n:ynat. n = 0 ∨ ∃m. n = ⫯m.
40 *
41 [ * /2 width=1 by or_introl/
42   #n @or_intror @(ex_intro … n) // (**) (* explicit constructor *)
43 | @or_intror @(ex_intro … (∞)) // (**) (* explicit constructor *)
44 ]
45 qed-.
46
47 lemma ysucc_iter_Y: ∀m. ysucc^m (∞) = ∞.
48 #m elim m -m //
49 #m #IHm whd in ⊢ (??%?); >IHm //
50 qed.
51
52 (* Inversion lemmas *********************************************************)
53
54 lemma ysucc_inv_inj: ∀m,n. ⫯m = ⫯n → m = n.
55 #m #n #H <(ypred_succ m) <(ypred_succ n) //
56 qed-.
57
58 lemma ysucc_inv_refl: ∀m. ⫯m = m → m = ∞.
59 * //
60 #m #H lapply (yinj_inj … H) -H (**) (* destruct lemma needed *)
61 #H elim (lt_refl_false m) //
62 qed-.
63
64 lemma ysucc_inv_inj_sn: ∀m2,n1. yinj m2 = ⫯n1 →
65                         ∃∃m1. n1 = yinj m1 & m2 = S m1.
66 #m2 * normalize
67 [ #n1 #H destruct /2 width=3 by ex2_intro/
68 | #H destruct
69 ]
70 qed-.
71
72 lemma ysucc_inv_inj_dx: ∀m2,n1. ⫯n1 = yinj m2  →
73                         ∃∃m1. n1 = yinj m1 & m2 = S m1.
74 /2 width=1 by ysucc_inv_inj_sn/ qed-.
75
76 lemma ysucc_inv_Y_sn: ∀m. ∞ = ⫯m → m = ∞.
77 * // normalize
78 #m #H destruct
79 qed-.
80
81 lemma ysucc_inv_Y_dx: ∀m. ⫯m = ∞ → m = ∞.
82 /2 width=1 by ysucc_inv_Y_sn/ qed-.
83
84 lemma ysucc_inv_O_sn: ∀m. yinj 0 = ⫯m → ⊥. (**) (* explicit coercion *)
85 #m #H elim (ysucc_inv_inj_sn … H) -H
86 #n #_ #H destruct
87 qed-.
88
89 lemma ysucc_inv_O_dx: ∀m. ⫯m = 0 → ⊥.
90 /2 width=2 by ysucc_inv_O_sn/ qed-.
91
92 (* Eliminators **************************************************************)
93
94 lemma ynat_ind: ∀R:predicate ynat.
95                 R 0 → (∀n:nat. R n → R (⫯n)) → R (∞) →
96                 ∀x. R x.
97 #R #H1 #H2 #H3 * // #n elim n -n /2 width=1 by/
98 qed-.