]> matita.cs.unibo.it Git - helm.git/blob - ground.ma
87d8d9baf4883976cffbbdccf916e3da2e2a6f9c
[helm.git] / 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/xoa_defs.ma".
14 include "lambda-delta/xoa_notation.ma".
15 include "lambda-delta/notation.ma".
16
17 (* ARITHMETICAL PROPERTIES **************************************************)
18
19 lemma plus_plus_comm_23: ∀m,n,p. m + n + p = m + p + n.
20 // qed.
21
22 lemma minus_le: ∀m,n. m - n ≤ m.
23 /2/ qed.
24
25 lemma plus_plus_minus_m_m: ∀e1,e2,d. e1 ≤ e2 → d + e1 + (e2 - e1) = d + e2.
26 /2/ qed.
27
28 lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → (m + n) - p = (m - p) + n.
29 #n #m #p #lepm @plus_to_minus <associative_plus
30 >(commutative_plus p) <plus_minus_m_m //
31 qed.
32
33 lemma lt_or_ge: ∀m,n. m < n ∨ n ≤ m.
34 #m #n elim (decidable_lt m n) /3/
35 qed.
36
37 lemma lt_false: ∀n. n < n → False.
38 #n #H lapply (lt_to_not_eq … H) -H #H elim H -H /2/
39 qed.
40
41 lemma plus_lt_false: ∀m,n. m + n < m → False.
42 #m #n #H1 lapply (le_plus_n_r n m) #H2
43 lapply (le_to_lt_to_lt … H2 H1) -H2 H1 #H
44 elim (lt_false … H)
45 qed.
46
47 lemma arith1: ∀n,h,m,p. n + h + m ≤ p + h → n + m ≤ p.
48 /2/ qed.
49
50 lemma arith2: ∀j,i,e,d. d + e ≤ i → d ≤ i - e + j.
51 /3/ qed.
52
53 lemma arith3: ∀m,n,p. p ≤ m → m + n - (m - p + n) = p.
54 /3/ qed.
55
56 lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2.
57 /2/ qed.
58
59 lemma arith5: ∀i,h,d. i + h ≤ d → d - i - h + (i + h) = d.
60 /2/ qed.