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 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 include "nat/lt_arith.ma".
19 called lt_mod_m_m, is in div_and_mod.
20 Other inequalities are also in lt_arith.ma.
23 theorem lt_div_S: \forall n,m. O < m \to
26 change with (n < m +(n/m)*m).
28 rewrite > (div_mod n m H) in ⊢ (? % ?).
34 theorem le_div: \forall n,m. O < n \to m/n \le m.
36 rewrite > (div_mod m n) in \vdash (? ? %)
37 [apply (trans_le ? (m/n*n))
38 [rewrite > times_n_SO in \vdash (? % ?).
40 [apply le_n|assumption]
47 theorem le_plus_mod: \forall m,n,q. O < q \to
48 (m+n) \mod q \le m \mod q + n \mod q .
50 elim (decidable_le q (m \mod q + n \mod q))
51 [apply not_lt_to_le.intro.
52 apply (le_to_not_lt q q)
54 |apply (le_to_lt_to_lt ? (m\mod q+n\mod q))
56 |apply (trans_lt ? ((m+n)\mod q))
58 |apply lt_mod_m_m.assumption
62 |cut ((m+n)\mod q = m\mod q+n\mod q)
63 [rewrite < Hcut.apply le_n
64 |apply (div_mod_spec_to_eq2 (m+n) q ((m+n)/q) ((m+n) \mod q) (m/q + n/q))
65 [apply div_mod_spec_div_mod.
67 |apply div_mod_spec_intro
68 [apply not_le_to_lt.assumption
69 |rewrite > (div_mod n q H) in ⊢ (? ? (? ? %) ?).
71 rewrite < assoc_plus in ⊢ (? ? ? %).
73 [rewrite > (div_mod m q) in ⊢ (? ? (? % ?) ?)
74 [rewrite > sym_times in ⊢ (? ? ? (? % ?)).
75 rewrite > distr_times_plus.
76 rewrite > sym_times in ⊢ (? ? ? (? (? % ?) ?)).
78 rewrite > assoc_plus in ⊢ (? ? ? %).
93 theorem le_plus_div: \forall m,n,q. O < q \to
94 m/q + n/q \le (m+n)/q.
96 apply (le_times_to_le q)
98 |rewrite > distr_times_plus.
100 rewrite > sym_times in ⊢ (? (? ? %) ?).
101 rewrite > sym_times in ⊢ (? ? %).
102 apply (le_plus_to_le ((m+n) \mod q)).
103 rewrite > sym_plus in ⊢ (? ? %).
104 rewrite < (div_mod ? ? H).
105 rewrite > (div_mod n q H) in ⊢ (? ? (? ? %)).
106 rewrite < assoc_plus.
107 rewrite > sym_plus in ⊢ (? ? (? ? %)).
108 rewrite < assoc_plus in ⊢ (? ? %).
110 rewrite > (div_mod m q H) in ⊢ (? ? (? % ?)).
111 rewrite > assoc_plus.
119 theorem le_times_to_le_div: \forall a,b,c:nat.
120 O \lt b \to (b*c) \le a \to c \le (a /b).
123 apply (lt_times_n_to_lt b)
125 |rewrite > sym_times.
126 apply (le_to_lt_to_lt ? a)
130 rewrite > (div_mod a b) in ⊢ (? % ?)
140 theorem le_times_to_le_div2: \forall m,n,q. O < q \to
141 n \le m*q \to n/q \le m.
143 apply (le_times_to_le q ? ? H).
145 rewrite > sym_times in ⊢ (? ? %).
146 apply (le_plus_to_le (n \mod q)).
149 [apply (trans_le ? (m*q))
158 theorem lt_m_nm: \forall n,m. O < m \to S O < n \to
162 [simplify.rewrite < plus_n_O.
163 rewrite > plus_n_O in ⊢ (? % ?).
164 apply lt_plus_r.assumption
166 rewrite > plus_n_O in ⊢ (? % ?).
175 theorem lt_times_to_lt: \forall i,n,m. O < i \to
176 i * n < i * m \to n < m.
178 apply not_le_to_lt.intro.
179 apply (lt_to_not_le ? ? H1).
184 theorem lt_times_to_lt_div: \forall m,n,q. n < m*q \to n/q < m.
186 apply (lt_times_to_lt q ? ? (lt_times_to_lt_O ? ? ? H)).
188 rewrite > sym_times in ⊢ (? ? %).
189 apply (le_plus_to_le (n \mod q)).
193 [apply (trans_le ? (m*q))
197 |apply (lt_times_to_lt_O ? ? ? H)
201 theorem lt_div: \forall n,m. O < m \to S O < n \to m/n < m.
203 apply lt_times_to_lt_div.
205 apply lt_m_nm;assumption.
208 theorem le_div_plus_S: \forall m,n,q. O < q \to
209 (m+n)/q \le S(m/q + n/q).
212 apply lt_times_to_lt_div.
213 change in ⊢ (? ? %) with (q + (q + (m/q+n/q)*q)).
215 rewrite > distr_times_plus.
217 rewrite < assoc_plus in ⊢ (? ? (? ? %)).
218 rewrite < assoc_plus.
219 rewrite > sym_plus in ⊢ (? ? (? % ?)).
220 rewrite > assoc_plus.
222 [change with (m < S(m/q)*q).
225 |rewrite > sym_times.
226 change with (n < S(n/q)*q).
232 theorem le_div_S_S_div: \forall n,m. O < m \to
233 (S n)/m \le S (n /m).
235 apply le_times_to_le_div2
238 rewrite > (div_mod n m H) in ⊢ (? (? %) ?).
247 theorem le_times_div_div_times: \forall a,n,m.O < m \to
250 apply le_times_to_le_div
252 |rewrite > sym_times.
253 rewrite > assoc_times.
255 rewrite > (div_mod n m H) in ⊢ (? ? %).
260 theorem monotonic_div: \forall n.O < n \to
261 monotonic nat le (\lambda m.div m n).
262 unfold monotonic.simplify.intros.
263 apply le_times_to_le_div
265 |apply (trans_le ? x)
266 [rewrite > sym_times.
267 rewrite > (div_mod x n H) in ⊢ (? ? %).
274 theorem le_div_times_m: \forall a,i,m. O < i \to O < m \to
275 (a * (m / i)) / m \le a / i.
277 apply (trans_le ? ((a*m/i)/m))
280 |apply le_times_div_div_times.
283 |rewrite > eq_div_div_div_times
284 [rewrite > sym_times in ⊢ (? (? ? %) ?).
285 rewrite < eq_div_div_div_times
288 |rewrite > lt_O_to_div_times
302 theorem le_div_times_Sm: \forall a,i,m. O < i \to O < m \to
303 a / i \le (a * S (m / i))/m.
305 apply (trans_le ? ((a * S (m / i))/((S (m/i))*i)))
306 [rewrite < (eq_div_div_div_times ? i)
307 [rewrite > lt_O_to_div_times
314 |apply le_times_to_le_div
316 |apply (trans_le ? (m*(a*S (m/i))/(S (m/i)*i)))
317 [apply le_times_div_div_times.
318 rewrite > (times_n_O O).
323 |rewrite > sym_times.
324 apply le_times_to_le_div2
325 [rewrite > (times_n_O O).