]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/ground.ma
d39b3faf46667050c302cd13a535595b9e383a21
[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/xoa_defs.ma".
14 include "lambda-delta/xoa_notation.ma".
15
16 (* ARITHMETICAL PROPERTIES **************************************************)
17
18 lemma plus_plus_comm_23: ∀m,n,p. m + n + p = m + p + n.
19 // qed.
20
21 lemma minus_plus_comm: ∀a,b,c. a - b - c = a - (c + b).
22 // qed.
23
24 lemma minus_le: ∀m,n. m - n ≤ m.
25 /2/ qed.
26
27 lemma plus_plus_minus_m_m: ∀e1,e2,d. e1 ≤ e2 → d + e1 + (e2 - e1) = d + e2.
28 /2/ qed.
29
30 lemma le_O_to_eq_O: ∀n. n ≤ 0 → n = 0.
31 /2/ qed.
32
33 lemma plus_le_minus: ∀a,b,c. a + b ≤ c → a ≤ c - b.
34 /2/ qed.
35
36 lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → (m + n) - p = (m - p) + n.
37 #n #m #p #lepm @plus_to_minus <associative_plus
38 >(commutative_plus p) <plus_minus_m_m //
39 qed.
40
41 lemma le_plus_minus: ∀a,b,c. c ≤ b → a + b - c = a + (b - c).
42 /2/ qed.
43
44 lemma minus_le_minus_minus_comm: ∀m,p,n. 
45                                  p ≤ m → m - p ≤ n → n + p - m = n - (m - p).
46 #m elim m -m
47 [ #p #n #H >(le_O_to_eq_O … H) -H //
48 | #m #IHm #p elim p -p //
49   #p #_ #n #Hpm <plus_n_Sm @IHm /2/
50 ]
51 qed.
52
53 lemma lt_refl_false: ∀n. n < n → False.
54 #n #H elim (lt_to_not_eq … H) -H /2/
55 qed.
56
57 lemma lt_zero_false: ∀n. n < 0 → False.
58 #n #H elim (lt_to_not_le … H) -H /2/
59 qed.
60
61 lemma lt_or_ge: ∀m,n. m < n ∨ n ≤ m.
62 #m #n elim (decidable_lt m n) /3/
63 qed.
64
65 lemma plus_lt_false: ∀m,n. m + n < m → False.
66 #m #n #H1 lapply (le_plus_n_r n m) #H2
67 lapply (le_to_lt_to_lt … H2 H1) -H2 H1 #H
68 elim (lt_refl_false … H)
69 qed.
70
71 lemma plus_S_eq_O_false: ∀n,m. n + S m = 0 → False.
72 #n #m <plus_n_Sm #H destruct
73 qed. 
74
75 lemma minus_minus_comm: ∀a,b,c. a - b - c = a - c - b.
76 /3/ qed.
77
78 lemma arith1: ∀n,h,m,p. n + h + m ≤ p + h → n + m ≤ p.
79 /2/ qed.
80
81 lemma arith6: ∀m,n. m < n → n - (n - m - 1) = m + 1.
82 #m #n #H >minus_plus <minus_minus //
83 qed.
84
85 lemma arith4: ∀h,d,e1,e2. d ≤ e1 + e2 → d + h ≤ e1 + h + e2.
86 /2/ qed.
87
88 lemma arith5: ∀i,h,d. i + h ≤ d → d - i - h + (i + h) = d.
89 /2/ qed.
90
91 lemma arith7: ∀i,d. i ≤ d → d - i + i = d.
92 /2/ qed.
93
94 lemma arith2: ∀j,i,e,d. d + e ≤ i → d ≤ i - e + j.
95 /3/ qed.
96
97 lemma arith3: ∀m,n,p. p ≤ m → m + n - (m - p + n) = p.
98 /3/ qed.
99