]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma
a29899d451a6b8b10152012ced8c0b7ed252839b
[helm.git] / matita / matita / contribs / lambdadelta / ground_2 / lib / arith.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/uparrow_1.ma".
16 include "ground_2/notation/functions/downarrow_1.ma".
17 include "arithmetics/nat.ma".
18 include "ground_2/lib/relations.ma".
19
20 (* ARITHMETICAL PROPERTIES **************************************************)
21
22 interpretation "nat successor" 'UpArrow m = (S m).
23
24 interpretation "nat predecessor" 'DownArrow m = (pred m).
25
26 interpretation "nat min" 'and x y = (min x y).
27
28 interpretation "nat max" 'or x y = (max x y).
29
30 (* Iota equations ***********************************************************)
31
32 lemma pred_O: pred 0 = 0.
33 normalize // qed.
34
35 lemma pred_S: ∀m. pred (S m) = m.
36 // qed.
37
38 lemma plus_S1: ∀x,y. ↑(x+y) = (↑x) + y.
39 // qed.
40
41 lemma max_O1: ∀n. n = (0 ∨ n).
42 // qed.
43
44 lemma max_O2: ∀n. n = (n ∨ 0).
45 // qed.
46
47 lemma max_SS: ∀n1,n2. ↑(n1∨n2) = (↑n1 ∨ ↑n2).
48 #n1 #n2 elim (decidable_le n1 n2) #H normalize
49 [ >(le_to_leb_true … H) | >(not_le_to_leb_false … H) ] -H //
50 qed.
51
52 (* Equalities ***************************************************************)
53
54 lemma plus_SO: ∀n. n + 1 = ↑n.
55 // qed.
56
57 lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.
58 // qed-.
59
60 lemma plus_to_minus_2: ∀m1,m2,n1,n2. n1 ≤ m1 → n2 ≤ m2 →
61                        m1+n2 = m2+n1 → m1-n1 = m2-n2.
62 #m1 #m2 #n1 #n2 #H1 #H2 #H
63 @plus_to_minus >plus_minus_associative //
64 qed-.
65
66 (* Note: uses minus_minus_comm, minus_plus_m_m, commutative_plus, plus_minus *)
67 lemma plus_minus_minus_be: ∀x,y,z. y ≤ z → z ≤ x → (x - z) + (z - y) = x - y.
68 #x #z #y #Hzy #Hyx >plus_minus // >commutative_plus >plus_minus //
69 qed-.
70
71 lemma lt_succ_pred: ∀m,n. n < m → m = ↑↓m.
72 #m #n #Hm >S_pred /2 width=2 by ltn_to_ltO/
73 qed-.
74
75 fact plus_minus_minus_be_aux: ∀i,x,y,z. y ≤ z → z ≤ x → i = z - y → x - z + i = x - y.
76 /2 width=1 by plus_minus_minus_be/ qed-.
77
78 lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p).
79 /2 by plus_minus/ qed-.
80
81 lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n.
82 /2 by plus_minus/ qed-.
83
84 lemma minus_minus_comm3: ∀n,x,y,z. n-x-y-z = n-y-z-x.
85 // qed.
86
87 lemma idempotent_max: ∀n:nat. n = (n ∨ n).
88 #n normalize >le_to_leb_true //
89 qed.
90
91 lemma associative_max: associative … max.
92 #x #y #z normalize
93 @(leb_elim x y) normalize #Hxy
94 @(leb_elim y z) normalize #Hyz //
95 [1,2: >le_to_leb_true /2 width=3 by transitive_le/
96 | >not_le_to_leb_false /4 width=3 by lt_to_not_le, not_le_to_lt, transitive_lt/
97   >not_le_to_leb_false //
98 ]
99 qed.
100
101 (* Properties ***************************************************************)
102
103 lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2).
104 #n1 elim n1 -n1 [| #n1 #IHn1 ] * [2,4: #n2 ]
105 [1,4: @or_intror #H destruct
106 | elim (IHn1 n2) -IHn1 /3 width=1 by or_intror, or_introl/
107 | /2 width=1 by or_introl/
108 ]
109 qed-.
110
111 lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m.
112 #m #n elim (lt_or_ge m n) /2 width=1 by or3_intro0/
113 #H elim H -m /2 width=1 by or3_intro1/
114 #m #Hm * /3 width=1 by not_le_to_lt, le_S_S, or3_intro2/
115 qed-.
116
117 lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y - z.
118 /3 width=1 by monotonic_le_minus_l/ qed.
119
120 lemma minus_le_trans_sn: ∀x1,x2. x1 ≤ x2 → ∀x. x1-x ≤ x2.
121 /2 width=3 by transitive_le/ qed.
122
123 lemma le_plus_to_minus_l: ∀a,b,c. a + b ≤ c → b ≤ c-a.
124 /2 width=1 by le_plus_to_minus_r/
125 qed-.
126
127 lemma le_plus_to_minus_comm: ∀n,m,p. n ≤ p+m → n-p ≤ m.
128 /2 width=1 by le_plus_to_minus/ qed-.
129
130 (* Note: this might interfere with nat.ma *)
131 lemma monotonic_lt_pred: ∀m,n. m < n → 0 < m → pred m < pred n.
132 #m #n #Hmn #Hm whd >(S_pred … Hm)
133 @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/
134 qed.
135
136 lemma lt_S_S: ∀x,y. x < y → ↑x < ↑y.
137 /2 width=1 by le_S_S/ qed.
138
139 lemma lt_S: ∀n,m. n < m → n < ↑m.
140 /2 width=1 by le_S/ qed.
141
142 lemma max_S1_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (↑n1 ∨ n2) ≤ ↑n.
143 /4 width=2 by to_max, le_maxr, le_S_S, le_S/ qed-.
144
145 lemma max_S2_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (n1 ∨ ↑n2) ≤ ↑n.
146 /2 width=1 by max_S1_le_S/ qed-.
147
148 (* Inversion & forward lemmas ***********************************************)
149
150 lemma lt_refl_false: ∀n. n < n → ⊥.
151 #n #H elim (lt_to_not_eq … H) -H /2 width=1 by/
152 qed-.
153
154 lemma lt_zero_false: ∀n. n < 0 → ⊥.
155 #n #H elim (lt_to_not_le … H) -H /2 width=1 by/
156 qed-.
157
158 lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥.
159 /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-.
160
161 lemma succ_inv_refl_sn: ∀x. ↑x = x → ⊥.
162 #x #H @(lt_le_false x (↑x)) //
163 qed-.
164
165 lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥.
166 #x #y #H lapply (le_n_O_to_eq … H) -H <plus_n_Sm #H destruct
167 qed-.
168
169 lemma le_plus_xySz_x_false: ∀y,z,x. x + y + S z ≤ x → ⊥.
170 #y #z #x elim x -x /3 width=1 by le_S_S_to_le/
171 #H elim (le_plus_xSy_O_false … H)
172 qed-.
173
174 lemma plus_xySz_x_false: ∀z,x,y. x + y + S z = x → ⊥.
175 /2 width=4 by le_plus_xySz_x_false/ qed-.
176
177 lemma plus_xSy_x_false: ∀y,x. x + S y = x → ⊥.
178 /2 width=4 by plus_xySz_x_false/ qed-.
179
180 lemma pred_inv_fix_sn: ∀x. ↓x = x → 0 = x.
181 * // #x <pred_Sn #H
182 elim (succ_inv_refl_sn x) //
183 qed-.
184
185 lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0.
186 // qed-.
187
188 lemma discr_plus_x_xy: ∀x,y. x = x + y → y = 0.
189 /2 width=2 by le_plus_minus_comm/ qed-.
190
191 lemma plus2_inv_le_sn: ∀m1,m2,n1,n2. m1 + n1 = m2 + n2 → m1 ≤ m2 → n2 ≤ n1.
192 #m1 #m2 #n1 #n2 #H #Hm
193 lapply (monotonic_le_plus_l n1 … Hm) -Hm >H -H
194 /2 width=2 by le_plus_to_le/
195 qed-.
196
197 lemma lt_S_S_to_lt: ∀x,y. ↑x < ↑y → x < y.
198 /2 width=1 by le_S_S_to_le/ qed-.
199
200 (* Note this should go in nat.ma *)
201 lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0.
202 #x @(nat_ind_plus … x) -x /2 width=1 by or_introl/
203 #x #_ #y @(nat_ind_plus … y) -y /2 width=1 by or_intror/
204 #y #_ >minus_plus_plus_l
205 #H lapply (discr_plus_xy_minus_xz … H) -H
206 #H destruct
207 qed-.
208
209 lemma lt_inv_O1: ∀n. 0 < n → ∃m. ↑m = n.
210 * /2 width=2 by ex_intro/
211 #H cases (lt_le_false … H) -H //
212 qed-.
213
214 lemma lt_inv_S1: ∀m,n. ↑m < n → ∃∃p. m < p & ↑p = n.
215 #m * /3 width=3 by lt_S_S_to_lt, ex2_intro/
216 #H cases (lt_le_false … H) -H //
217 qed-.
218
219 lemma lt_inv_gen: ∀y,x. x < y → ∃∃z. x ≤ z & ↑z = y.
220 * /3 width=3 by le_S_S_to_le, ex2_intro/
221 #x #H elim (lt_le_false … H) -H //
222 qed-.
223
224 lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0.
225 /2 width=1 by plus_le_0/ qed-.
226
227 lemma plus_inv_S3_sn: ∀x1,x2,x3. x1+x2 = ↑x3 →
228                       ∨∨ ∧∧ x1 = 0 & x2 = ↑x3
229                        | ∃∃y1. x1 = ↑y1 & y1 + x2 = x3.
230 * /3 width=1 by or_introl, conj/
231 #x1 #x2 #x3 <plus_S1 #H destruct
232 /3 width=3 by ex2_intro, or_intror/
233 qed-.
234
235 lemma plus_inv_S3_dx: ∀x2,x1,x3. x1+x2 = ↑x3 →
236                       ∨∨ ∧∧ x2 = 0 & x1 = ↑x3
237                        | ∃∃y2. x2 = ↑y2 & x1 + y2 = x3.
238 * /3 width=1 by or_introl, conj/
239 #x2 #x1 #x3 <plus_n_Sm #H destruct
240 /3 width=3 by ex2_intro, or_intror/
241 qed-.
242
243 lemma max_inv_O3: ∀x,y. (x ∨ y) = 0 → 0 = x ∧ 0 = y.
244 /4 width=2 by le_maxr, le_maxl, le_n_O_to_eq, conj/
245 qed-.
246
247 lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
248 * /2 width=1 by conj/ #x #y normalize #H destruct
249 qed-.
250
251 lemma nat_split: ∀x. x = 0 ∨ ∃y. ↑y = x.
252 * /3 width=2 by ex_intro, or_introl, or_intror/
253 qed-.
254
255 lemma lt_elim: ∀R:relation nat.
256                (∀n2. R O (↑n2)) →
257                (∀n1,n2. R n1 n2 → R (↑n1) (↑n2)) →
258                ∀n2,n1. n1 < n2 → R n1 n2.
259 #R #IH1 #IH2 #n2 elim n2 -n2
260 [ #n1 #H elim (lt_le_false … H) -H //
261 | #n2 #IH * /4 width=1 by lt_S_S_to_lt/
262 ]
263 qed-.
264
265 lemma le_elim: ∀R:relation nat.
266                (∀n2. R O (n2)) →
267                (∀n1,n2. R n1 n2 → R (↑n1) (↑n2)) →
268                ∀n1,n2. n1 ≤ n2 → R n1 n2.
269 #R #IH1 #IH2 #n1 #n2 @(nat_elim2 … n1 n2) -n1 -n2
270 /4 width=1 by monotonic_pred/ -IH1 -IH2
271 #n1 #H elim (lt_le_false … H) -H //
272 qed-.
273
274 (* Iterators ****************************************************************)
275
276 (* Note: see also: lib/arithemetics/bigops.ma *)
277 rec definition iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
278   match n with
279    [ O   ⇒ nil
280    | S k ⇒ op (iter k B op nil)
281    ].
282
283 interpretation "iterated function" 'exp op n = (iter n ? op).
284
285 lemma iter_O: ∀B:Type[0]. ∀f:B→B.∀b. f^0 b = b.
286 // qed.
287
288 lemma iter_S: ∀B:Type[0]. ∀f:B→B.∀b,l. f^(S l) b = f (f^l b).
289 // qed.
290
291 lemma iter_n_Sm: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^l (f b) = f (f^l b).
292 #B #f #b #l elim l -l normalize //
293 qed.
294
295 lemma iter_plus: ∀B:Type[0]. ∀f:B→B. ∀b,l1,l2. f^(l1+l2) b = f^l1 (f^l2 b).
296 #B #f #b #l1 elim l1 -l1 normalize //
297 qed.
298
299 (* Trichotomy operator ******************************************************)
300
301 (* Note: this is "if eqb n1 n2 then a2 else if leb n1 n2 then a1 else a3" *)
302 rec definition tri (A:Type[0]) n1 n2 a1 a2 a3 on n1 : A ≝
303   match n1 with
304   [ O    ⇒ match n2 with [ O ⇒ a2 | S n2 ⇒ a1 ]
305   | S n1 ⇒ match n2 with [ O ⇒ a3 | S n2 ⇒ tri A n1 n2 a1 a2 a3 ]
306   ].
307
308 lemma tri_lt: ∀A,a1,a2,a3,n2,n1. n1 < n2 → tri A n1 n2 a1 a2 a3 = a1.
309 #A #a1 #a2 #a3 #n2 elim n2 -n2
310 [ #n1 #H elim (lt_zero_false … H)
311 | #n2 #IH #n1 elim n1 -n1 /3 width=1 by monotonic_lt_pred/
312 ]
313 qed.
314
315 lemma tri_eq: ∀A,a1,a2,a3,n. tri A n n a1 a2 a3 = a2.
316 #A #a1 #a2 #a3 #n elim n -n normalize //
317 qed.
318
319 lemma tri_gt: ∀A,a1,a2,a3,n1,n2. n2 < n1 → tri A n1 n2 a1 a2 a3 = a3.
320 #A #a1 #a2 #a3 #n1 elim n1 -n1
321 [ #n2 #H elim (lt_zero_false … H)
322 | #n1 #IH #n2 elim n2 -n2 /3 width=1 by monotonic_lt_pred/
323 ]
324 qed.