]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma
- improved arithmetics for natural numbers with infinity
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / ynat / ynat_plus.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/ynat/ynat_minus.ma".
16
17 (* NATURAL NUMBERS WITH INFINITY ********************************************)
18
19 (* addition *)
20 definition yplus: ynat → ynat → ynat ≝ λx,y. match y with
21 [ yinj n ⇒ ysucc^n x
22 | Y      ⇒ Y
23 ].
24
25 interpretation "ynat plus" 'plus x y = (yplus x y).
26
27 (* Properties on successor **************************************************)
28
29 lemma yplus_succ2: ∀m,n. m + ⫯n = ⫯(m + n).
30 #m * //
31 qed.
32
33 lemma yplus_succ1: ∀m,n. ⫯m + n = ⫯(m + n).
34 #m * normalize //
35 qed.
36
37 lemma yplus_succ_swap: ∀m,n. m + ⫯n = ⫯m + n.
38 // qed.
39
40 lemma yplus_SO2: ∀m. m + 1 = ⫯m.
41 * //
42 qed.
43
44 (* Basic properties *********************************************************)
45
46 lemma yplus_inj: ∀n,m. yinj m + yinj n = yinj (m + n).
47 #n elim n -n [ normalize // ]
48 #n #IHn #m >(yplus_succ2 ? n) >IHn -IHn
49 <plus_n_Sm //
50 qed.
51
52 lemma yplus_Y1: ∀m. ∞ + m = ∞.
53 * normalize //
54 qed.
55
56 lemma yplus_comm: commutative … yplus.
57 * [ #m ] * [1,3: #n ] //
58 normalize >ysucc_iter_Y //
59 qed.
60
61 lemma yplus_assoc: associative … yplus.
62 #x #y * // #z cases y -y
63 [ #y >yplus_inj whd in ⊢ (??%%); <iter_plus //
64 | >yplus_Y1 //
65 ]
66 qed.
67
68 lemma yplus_O_sn: ∀n:ynat. 0 + n = n.
69 #n >yplus_comm // qed.
70
71 (* Basic inversion lemmas ***************************************************)
72
73 lemma yplus_inv_inj: ∀z,y,x. x + y = yinj z →
74                      ∃∃m,n. m + n = z & x = yinj m & y = yinj n.
75 #z * [2: normalize #x #H destruct ]
76 #y * [2: >yplus_Y1 #H destruct ]
77 /3 width=5 by yinj_inj, ex3_2_intro/
78 qed-.
79
80 (* Properties on order ******************************************************)
81
82 lemma yle_plus_dx2: ∀n,m. n ≤ m + n.
83 * //
84 #n elim n -n //
85 #n #IHn #m >(yplus_succ2 ? n) @(yle_succ n) // (**) (* full auto fails *)
86 qed.
87
88 lemma yle_plus_dx1: ∀n,m. m ≤ m + n.
89 // qed.
90
91 (* Forward lemmas on order **************************************************)
92
93 lemma yle_fwd_plus_sn2: ∀x,y,z. x + y ≤ z → y ≤ z.
94 /2 width=3 by yle_trans/ qed-.
95
96 lemma yle_fwd_plus_sn1: ∀x,y,z. x + y ≤ z → x ≤ z.
97 /2 width=3 by yle_trans/ qed-.