]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/ground.ma
ground: some arithmetical properties added
[helm.git] / matita / matita / lib / lambda-delta / ground.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic
3     ||A||  Library of Mathematics, developed at the Computer Science
4     ||T||  Department of the University of Bologna, Italy.
5     ||I||
6     ||T||
7     ||A||  This file is distributed under the terms of the
8     \   /  GNU General Public License Version 2
9      \ /
10       V_______________________________________________________________ *)
11
12 include "basics/list.ma".
13 include "lambda-delta/notation.ma".
14
15 (* ARITHMETICAL PROPERTIES **************************************************)
16
17 lemma plus_plus_comm_23: ∀m,n,p. m + n + p = m + p + n.
18 // qed.
19
20 lemma minus_le: ∀m,n. m - n ≤ m.
21 /2/ qed.
22
23
24 lemma plus_plus_minus_m_m: ∀e1,e2,d. e1 ≤ e2 → d + e1 + (e2 - e1) = d + e2.
25 /2/ qed.
26
27 lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → (m + n) - p = (m - p) + n.
28 #n #m #p #lepm @plus_to_minus <associative_plus
29 >(commutative_plus p) <plus_minus_m_m //
30 qed.
31
32 lemma lt_false: ∀n. n < n → False.
33 #n #H lapply (lt_to_not_eq … H) -H #H elim H -H /2/
34 qed.
35
36 lemma arith1: ∀n,h,m,p. n + h + m ≤ p + h → n + m ≤ p.
37 /2/ qed.
38
39 lemma arith2: ∀j,i,e,d. d + e ≤ i → d ≤ i - e + j.
40 /3/ qed.
41
42 lemma arith3: ∀m,n,p. p ≤ m → m + n - (m - p + n) = p.
43 /3/ qed.
44
45 lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2.
46 /2/ qed.