1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
16 include "nat/le_arith.ma".
17 include "nat/compare.ma".
19 let rec minus n m \def
25 | (S q) \Rightarrow minus p q ]].
27 (*CSC: the URI must disappear: there is a bug now *)
28 interpretation "natural minus" 'minus x y = (cic:/matita/nat/minus/minus.con x y).
30 theorem minus_n_O: \forall n:nat.n=n-O.
31 intros.elim n.simplify.reflexivity.
35 theorem minus_n_n: \forall n:nat.O=n-n.
36 intros.elim n.simplify.
41 theorem minus_Sn_n: \forall n:nat. S O = (S n)-n.
47 theorem minus_Sn_m: \forall n,m:nat. m \leq n \to (S n)-m = S (n-m).
50 (\lambda n,m.m \leq n \to (S n)-m = S (n-m))).
51 intros.apply (le_n_O_elim n1 H).
53 intros.simplify.reflexivity.
54 intros.rewrite < H.reflexivity.
55 apply le_S_S_to_le. assumption.
58 theorem eq_minus_S_pred: \forall n,m. n - (S m) = pred(n -m).
61 |intro.simplify.autobatch
62 |intros.simplify.assumption
67 \forall n,m,p:nat. m \leq n \to (n-m)+p = (n+p)-m.
70 (\lambda n,m.\forall p:nat.m \leq n \to (n-m)+p = (n+p)-m)).
71 intros.apply (le_n_O_elim ? H).
72 simplify.rewrite < minus_n_O.reflexivity.
73 intros.simplify.reflexivity.
74 intros.simplify.apply H.apply le_S_S_to_le.assumption.
77 theorem minus_plus_m_m: \forall n,m:nat.n = (n+m)-m.
80 rewrite < minus_n_O.apply plus_n_O.
84 change with (S n2 = (S n2 + n)-n).
88 theorem plus_minus_m_m: \forall n,m:nat.
89 m \leq n \to n = (n-m)+m.
91 apply (nat_elim2 (\lambda n,m.m \leq n \to n = (n-m)+m)).
92 intros.apply (le_n_O_elim n1 H).
94 intros.simplify.rewrite < plus_n_O.reflexivity.
95 intros.simplify.rewrite < sym_plus.simplify.
96 apply eq_f.rewrite < sym_plus.apply H.
97 apply le_S_S_to_le.assumption.
100 theorem minus_to_plus :\forall n,m,p:nat.m \leq n \to n-m = p \to
102 intros.apply (trans_eq ? ? ((n-m)+m)).
103 apply plus_minus_m_m.
108 theorem plus_to_minus :\forall n,m,p:nat.
111 apply (inj_plus_r m).
115 apply plus_minus_m_m.rewrite > H.
120 theorem minus_S_S : \forall n,m:nat.
121 eq nat (minus (S n) (S m)) (minus n m).
126 theorem minus_pred_pred : \forall n,m:nat. lt O n \to lt O m \to
127 eq nat (minus (pred n) (pred m)) (minus n m).
129 apply (lt_O_n_elim n H).intro.
130 apply (lt_O_n_elim m H1).intro.
131 simplify.reflexivity.
134 theorem eq_minus_n_m_O: \forall n,m:nat.
135 n \leq m \to n-m = O.
137 apply (nat_elim2 (\lambda n,m.n \leq m \to n-m = O)).
138 intros.simplify.reflexivity.
139 intros.apply False_ind.
143 simplify.apply H.apply le_S_S_to_le. apply H1.
146 theorem le_SO_minus: \forall n,m:nat.S n \leq m \to S O \leq m-n.
147 intros.elim H.elim (minus_Sn_n n).apply le_n.
148 rewrite > minus_Sn_m.
149 apply le_S.assumption.
150 apply lt_to_le.assumption.
153 theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)).
154 intros.apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n)))).
155 intro.elim n1.simplify.apply le_n_Sn.
156 simplify.rewrite < minus_n_O.apply le_n.
157 intros.simplify.apply le_n_Sn.
158 intros.simplify.apply H.
161 theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p.
162 intros 3.simplify.intro.
163 apply (trans_le (m-n) (S (m-(S n))) p).
164 apply minus_le_S_minus_S.
168 theorem le_minus_m: \forall n,m:nat. n-m \leq n.
169 intros.apply (nat_elim2 (\lambda m,n. n-m \leq n)).
170 intros.rewrite < minus_n_O.apply le_n.
171 intros.simplify.apply le_n.
172 intros.simplify.apply le_S.assumption.
175 theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n.
176 intros.apply (lt_O_n_elim n H).intro.
177 apply (lt_O_n_elim m H1).intro.
178 simplify.unfold lt.apply le_S_S.apply le_minus_m.
181 theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m.
183 apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m)).
185 simplify.intros. assumption.
186 simplify.intros.apply le_S_S.apply H.assumption.
190 theorem monotonic_le_minus_r:
191 \forall p,q,n:nat. q \leq p \to n-p \le n-q.
192 simplify.intros 2.apply (nat_elim2
193 (\lambda p,q.\forall a.q \leq p \to a-p \leq a-q)).
194 intros.apply (le_n_O_elim n H).apply le_n.
195 intros.rewrite < minus_n_O.
197 intros.elim a.simplify.apply le_n.
198 simplify.apply H.apply le_S_S_to_le.assumption.
201 theorem le_minus_to_plus: \forall n,m,p. (le (n-m) p) \to (le n (p+m)).
202 intros 2.apply (nat_elim2 (\lambda n,m.\forall p.(le (n-m) p) \to (le n (p+m)))).
204 simplify.intros.rewrite < plus_n_O.assumption.
207 apply le_S_S.apply H.
211 theorem le_plus_to_minus: \forall n,m,p. (le n (p+m)) \to (le (n-m) p).
212 intros 2.apply (nat_elim2 (\lambda n,m.\forall p.(le n (p+m)) \to (le (n-m) p))).
213 intros.simplify.apply le_O_n.
214 intros 2.rewrite < plus_n_O.intro.simplify.assumption.
215 intros.simplify.apply H.
216 apply le_S_S_to_le.rewrite > plus_n_Sm.assumption.
219 (* the converse of le_plus_to_minus does not hold *)
220 theorem le_plus_to_minus_r: \forall n,m,p. (le (n+m) p) \to (le n (p-m)).
221 intros 3.apply (nat_elim2 (\lambda m,p.(le (n+m) p) \to (le n (p-m)))).
222 intro.rewrite < plus_n_O.rewrite < minus_n_O.intro.assumption.
223 intro.intro.cut (n=O).rewrite > Hcut.apply le_O_n.
224 apply sym_eq. apply le_n_O_to_eq.
225 apply (trans_le ? (n+(S n1))).
227 apply le_plus_n.assumption.
229 apply H.apply le_S_S_to_le.
230 rewrite > plus_n_Sm.assumption.
233 (* minus and lt - to be completed *)
234 theorem lt_minus_l: \forall m,l,n:nat.
235 l < m \to m \le n \to n - m < n - l.
237 [intros.apply False_ind.apply (not_le_Sn_O ? H)
238 |intros.rewrite < minus_n_O.
241 generalize in match H2.
243 [intros.apply False_ind.apply (not_le_Sn_O ? H3)
249 |apply le_S_S_to_le.assumption
255 theorem lt_minus_r: \forall n,m,l:nat.
256 n \le l \to l < m \to l -n < m -n.
259 |rewrite > eq_minus_S_pred.
260 rewrite > eq_minus_S_pred.
262 [unfold lt.apply le_plus_to_minus_r.applyS H1
263 |apply H[autobatch|assumption]
268 lemma lt_to_lt_O_minus : \forall m,n.
271 unfold. apply le_plus_to_minus_r. unfold in H. rewrite > sym_plus.
277 theorem lt_minus_to_plus: \forall n,m,p. (lt n (p-m)) \to (lt (n+m) p).
278 intros 3.apply (nat_elim2 (\lambda m,p.(lt n (p-m)) \to (lt (n+m) p))).
279 intro.rewrite < plus_n_O.rewrite < minus_n_O.intro.assumption.
280 simplify.intros.apply False_ind.apply (not_le_Sn_O n H).
281 simplify.intros.unfold lt.
287 theorem lt_O_minus_to_lt: \forall a,b:nat.
288 O \lt b-a \to a \lt b.
290 rewrite > (plus_n_O a).
291 rewrite > (sym_plus a O).
292 apply (lt_minus_to_plus O a b).
296 theorem lt_minus_to_lt_plus:
297 \forall n,m,p. n - m < p \to n < m + p.
299 apply (nat_elim2 ? ? ? ? n m)
300 [simplify.intros.autobatch.
301 |intros 2.rewrite < minus_n_O.
313 theorem lt_plus_to_lt_minus:
314 \forall n,m,p. m \le n \to n < m + p \to n - m < p.
316 apply (nat_elim2 ? ? ? ? n m)
318 apply (le_n_O_elim ? H).
319 simplify.intros.assumption
320 |simplify.intros.assumption.
324 [apply le_S_S_to_le.assumption
325 |apply le_S_S_to_le.apply H2
330 theorem minus_m_minus_mn: \forall n,m. n\le m \to n=m-(m-n).
337 theorem distributive_times_minus: distributive nat times minus.
340 apply ((leb_elim z y)).
341 intro.cut (x*(y-z)+x*z = (x*y-x*z)+x*z).
342 apply (inj_plus_l (x*z)).assumption.
343 apply (trans_eq nat ? (x*y)).
344 rewrite < distr_times_plus.rewrite < (plus_minus_m_m ? ? H).reflexivity.
345 rewrite < plus_minus_m_m.
347 apply le_times_r.assumption.
348 intro.rewrite > eq_minus_n_m_O.
349 rewrite > (eq_minus_n_m_O (x*y)).
350 rewrite < sym_times.simplify.reflexivity.
351 apply le_times_r.apply lt_to_le.apply not_le_to_lt.assumption.
352 apply lt_to_le.apply not_le_to_lt.assumption.
355 theorem distr_times_minus: \forall n,m,p:nat. n*(m-p) = n*m-n*p
356 \def distributive_times_minus.
358 theorem eq_minus_plus_plus_minus: \forall n,m,p:nat. p \le m \to (n+m)-p = n+(m-p).
361 rewrite > sym_plus in \vdash (? ? ? %).
362 rewrite > assoc_plus.
363 rewrite < plus_minus_m_m.
364 reflexivity.assumption.
367 theorem eq_minus_minus_minus_plus: \forall n,m,p:nat. (n-m)-p = n-(m+p).
369 cut (m+p \le n \or m+p \nleq n).
371 symmetry.apply plus_to_minus.
372 rewrite > assoc_plus.rewrite > (sym_plus p).rewrite < plus_minus_m_m.
373 rewrite > sym_plus.rewrite < plus_minus_m_m.
375 apply (trans_le ? (m+p)).
376 rewrite < sym_plus.apply le_plus_n.
378 apply le_plus_to_minus_r.rewrite > sym_plus.assumption.
379 rewrite > (eq_minus_n_m_O n (m+p)).
380 rewrite > (eq_minus_n_m_O (n-m) p).
382 apply le_plus_to_minus.apply lt_to_le. rewrite < sym_plus.
383 apply not_le_to_lt. assumption.
384 apply lt_to_le.apply not_le_to_lt.assumption.
385 apply (decidable_le (m+p) n).
388 theorem eq_plus_minus_minus_minus: \forall n,m,p:nat. p \le m \to m \le n \to
393 rewrite < assoc_plus.
394 rewrite < plus_minus_m_m.
396 rewrite < plus_minus_m_m.reflexivity.
397 assumption.assumption.