]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma
- ng_kernel: catched Invalid_argument "List.nth" in typeof of Rel i (raied when i...
[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/successor_1.ma".
16 include "ground_2/notation/functions/predecessor_1.ma".
17 include "arithmetics/nat.ma".
18 include "ground_2/lib/star.ma".
19
20 (* ARITHMETICAL PROPERTIES **************************************************)
21
22 interpretation "nat successor" 'Successor m = (S m).
23
24 interpretation "nat predecessor" 'Predecessor m = (pred m).
25
26 (* Iota equations ***********************************************************)
27
28 lemma pred_O: pred 0 = 0.
29 normalize // qed.
30
31 lemma pred_S: ∀m. pred (S m) = m.
32 // qed.
33
34 (* Equations ****************************************************************)
35
36 lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m.
37 // qed-.
38
39 (* Note: uses minus_minus_comm, minus_plus_m_m, commutative_plus, plus_minus *)
40 lemma plus_minus_minus_be: ∀x,y,z. y ≤ z → z ≤ x → (x - z) + (z - y) = x - y.
41 #x #z #y #Hzy #Hyx >plus_minus // >commutative_plus >plus_minus //
42 qed-.
43
44 fact plus_minus_minus_be_aux: ∀i,x,y,z. y ≤ z → z ≤ x → i = z - y → x - z + i = x - y.
45 /2 width=1 by plus_minus_minus_be/ qed-.
46
47 lemma plus_n_2: ∀n. n + 2 = n + 1 + 1.
48 // qed.
49
50 lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p).
51 /2 by plus_minus/ qed.
52
53 lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n.
54 /2 by plus_minus/ qed.
55
56 lemma minus_minus_comm3: ∀n,x,y,z. n-x-y-z = n-y-z-x.
57 // qed.
58
59 lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b.
60 #a #b #c1 #H >minus_minus_comm >minus_le_minus_minus_comm //
61 qed.
62
63 lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b.
64 #a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1 by arith_b1/
65 qed.
66
67 lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b.
68 /3 by monotonic_le_minus_l, le_to_le_to_eq, le_n/ qed.
69
70 lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b →
71                 a1 - c1 + a2 - (b - c1) = a1 + a2 - b.
72 #a1 #a2 #b #c1 #H1 #H2 >plus_minus /2 width=1 by arith_b2/
73 qed.
74
75 lemma arith_i: ∀x,y,z. y < x → x+z-y-1 = x-y-1+z.
76 /2 width=1 by plus_minus/ qed-.
77
78 (* Properties ***************************************************************)
79
80 lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2).
81 #n1 elim n1 -n1 [| #n1 #IHn1 ] * [2,4: #n2 ]
82 [1,4: @or_intror #H destruct
83 | elim (IHn1 n2) -IHn1 /3 width=1 by or_intror, or_introl/
84 | /2 width=1 by or_introl/
85 ]
86 qed-.
87
88 lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m.
89 #m #n elim (lt_or_ge m n) /2 width=1 by or3_intro0/
90 #H elim H -m /2 width=1 by or3_intro1/
91 #m #Hm * /3 width=1 by not_le_to_lt, le_S_S, or3_intro2/
92 qed-.
93
94 fact le_repl_sn_conf_aux: ∀x,y,z:nat. x ≤ z → x = y → y ≤ z.
95 // qed-.
96
97 fact le_repl_sn_trans_aux: ∀x,y,z:nat. x ≤ z → y = x → y ≤ z.
98 // qed-.
99
100 lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y - z.
101 /3 width=1 by monotonic_le_minus_l/ qed.
102
103 (* Note: this might interfere with nat.ma *)
104 lemma monotonic_lt_pred: ∀m,n. m < n → O < m → pred m < pred n.
105 #m #n #Hmn #Hm whd >(S_pred … Hm)
106 @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/
107 qed.
108
109 lemma lt_S_S: ∀x,y. x < y → ⫯x < ⫯y.
110 /2 width=1 by le_S_S/ qed.
111
112 lemma arith_j: ∀x,y,z. x-y-1 ≤ x-(y-z)-1.
113 /3 width=1 by monotonic_le_minus_l, monotonic_le_minus_r/ qed.
114
115 lemma arith_k_sn: ∀z,x,y,n. z < x → x+n ≤ y → x-z-1+n ≤ y-z-1.
116 #z #x #y #n #Hzx #Hxny
117 >plus_minus [2: /2 width=1 by monotonic_le_minus_r/ ]
118 >plus_minus [2: /2 width=1 by lt_to_le/ ]
119 /2 width=1 by monotonic_le_minus_l2/
120 qed.
121
122 lemma arith_k_dx: ∀z,x,y,n. z < x → y ≤ x+n → y-z-1 ≤ x-z-1+n.
123 #z #x #y #n #Hzx #Hyxn
124 >plus_minus [2: /2 width=1 by monotonic_le_minus_r/ ]
125 >plus_minus [2: /2 width=1 by lt_to_le/ ]
126 /2 width=1 by monotonic_le_minus_l2/
127 qed.
128
129 (* Inversion & forward lemmas ***********************************************)
130
131 lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0.
132 // qed-.
133
134 lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y.
135 /2 width=1 by monotonic_pred/ qed-.
136
137 lemma lt_refl_false: ∀n. n < n → ⊥.
138 #n #H elim (lt_to_not_eq … H) -H /2 width=1 by/
139 qed-.
140
141 lemma lt_zero_false: ∀n. n < 0 → ⊥.
142 #n #H elim (lt_to_not_le … H) -H /2 width=1 by/
143 qed-.
144
145 lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥.
146 /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-.
147
148 lemma pred_inv_refl: ∀m. pred m = m → m = 0.
149 * // normalize #m #H elim (lt_refl_false m) //
150 qed-.
151
152 lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥.
153 #x #y #H lapply (le_n_O_to_eq … H) -H <plus_n_Sm #H destruct
154 qed-.
155
156 lemma le_plus_xySz_x_false: ∀y,z,x. x + y + S z ≤ x → ⊥.
157 #y #z #x elim x -x /3 width=1 by le_S_S_to_le/
158 #H elim (le_plus_xSy_O_false … H)
159 qed-.
160
161 lemma plus_xySz_x_false: ∀z,x,y. x + y + S z = x → ⊥.
162 /2 width=4 by le_plus_xySz_x_false/ qed-.
163
164 lemma plus_xSy_x_false: ∀y,x. x + S y = x → ⊥.
165 /2 width=4 by plus_xySz_x_false/ qed-.
166
167 (* Note this should go in nat.ma *)
168 lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0.
169 #x @(nat_ind_plus … x) -x /2 width=1 by or_introl/
170 #x #_ #y @(nat_ind_plus … y) -y /2 width=1 by or_intror/
171 #y #_ >minus_plus_plus_l
172 #H lapply (discr_plus_xy_minus_xz … H) -H
173 #H destruct
174 qed-.
175
176 lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y.
177 * /2 width=1 by conj/ #x #y normalize #H destruct
178 qed-.
179
180 lemma lt_S_S_to_lt: ∀x,y. ⫯x < ⫯y → x < y.
181 /2 width=1 by le_S_S_to_le/ qed-.
182
183 lemma lt_elim: ∀R:relation nat.
184                (∀n2. R O (⫯n2)) →
185                (∀n1,n2. R n1 n2 → R (⫯n1) (⫯n2)) →
186                ∀n2,n1. n1 < n2 → R n1 n2.
187 #R #IH1 #IH2 #n2 elim n2 -n2
188 [ #n1 #H elim (lt_le_false … H) -H //
189 | #n2 #IH * /4 width=1 by lt_S_S_to_lt/
190 ]
191 qed-.
192
193 (* Iterators ****************************************************************)
194
195 (* Note: see also: lib/arithemetics/bigops.ma *)
196 let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝
197   match n with
198    [ O   ⇒ nil
199    | S k ⇒ op (iter k B op nil)
200    ].
201
202 interpretation "iterated function" 'exp op n = (iter n ? op).
203
204 lemma iter_O: ∀B:Type[0]. ∀f:B→B.∀b. f^0 b = b.
205 // qed.
206
207 lemma iter_S: ∀B:Type[0]. ∀f:B→B.∀b,l. f^(S l) b = f (f^l b).
208 // qed.
209
210 lemma iter_SO: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^(l+1) b = f (f^l b).
211 #B #f #b #l >commutative_plus //
212 qed.
213
214 lemma iter_n_Sm: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^l (f b) = f (f^l b).
215 #B #f #b #l elim l -l normalize //
216 qed.
217
218 lemma iter_plus: ∀B:Type[0]. ∀f:B→B. ∀b,l1,l2. f^(l1+l2) b = f^l1 (f^l2 b).
219 #B #f #b #l1 elim l1 -l1 normalize //
220 qed.
221
222 (* Trichotomy operator ******************************************************)
223
224 (* Note: this is "if eqb n1 n2 then a2 else if leb n1 n2 then a1 else a3" *)
225 let rec tri (A:Type[0]) n1 n2 a1 a2 a3 on n1 : A ≝
226   match n1 with
227   [ O    ⇒ match n2 with [ O ⇒ a2 | S n2 ⇒ a1 ]
228   | S n1 ⇒ match n2 with [ O ⇒ a3 | S n2 ⇒ tri A n1 n2 a1 a2 a3 ]
229   ].
230
231 lemma tri_lt: ∀A,a1,a2,a3,n2,n1. n1 < n2 → tri A n1 n2 a1 a2 a3 = a1.
232 #A #a1 #a2 #a3 #n2 elim n2 -n2
233 [ #n1 #H elim (lt_zero_false … H)
234 | #n2 #IH #n1 elim n1 -n1 /3 width=1 by monotonic_lt_pred/
235 ]
236 qed.
237
238 lemma tri_eq: ∀A,a1,a2,a3,n. tri A n n a1 a2 a3 = a2.
239 #A #a1 #a2 #a3 #n elim n -n normalize //
240 qed.
241
242 lemma tri_gt: ∀A,a1,a2,a3,n1,n2. n2 < n1 → tri A n1 n2 a1 a2 a3 = a3.
243 #A #a1 #a2 #a3 #n1 elim n1 -n1
244 [ #n2 #H elim (lt_zero_false … H)
245 | #n1 #IH #n2 elim n2 -n2 /3 width=1 by monotonic_lt_pred/
246 ]
247 qed.