author | haftmann |
Fri, 27 Dec 2013 14:35:14 +0100 | |
changeset 54868 | bab6cade3cc5 |
parent 54489 | 03ff4d1e6784 |
child 58040 | 9a867afaab5a |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: HOL/Archimedean_Field.thy |
2 |
Author: Brian Huffman |
|
30096 | 3 |
*) |
4 |
||
5 |
header {* Archimedean Fields, Floor and Ceiling Functions *} |
|
6 |
||
7 |
theory Archimedean_Field |
|
8 |
imports Main |
|
9 |
begin |
|
10 |
||
11 |
subsection {* Class of Archimedean fields *} |
|
12 |
||
13 |
text {* Archimedean fields have no infinite elements. *} |
|
14 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
15 |
class archimedean_field = linordered_field + |
30096 | 16 |
assumes ex_le_of_int: "\<exists>z. x \<le> of_int z" |
17 |
||
18 |
lemma ex_less_of_int: |
|
19 |
fixes x :: "'a::archimedean_field" shows "\<exists>z. x < of_int z" |
|
20 |
proof - |
|
21 |
from ex_le_of_int obtain z where "x \<le> of_int z" .. |
|
22 |
then have "x < of_int (z + 1)" by simp |
|
23 |
then show ?thesis .. |
|
24 |
qed |
|
25 |
||
26 |
lemma ex_of_int_less: |
|
27 |
fixes x :: "'a::archimedean_field" shows "\<exists>z. of_int z < x" |
|
28 |
proof - |
|
29 |
from ex_less_of_int obtain z where "- x < of_int z" .. |
|
30 |
then have "of_int (- z) < x" by simp |
|
31 |
then show ?thesis .. |
|
32 |
qed |
|
33 |
||
34 |
lemma ex_less_of_nat: |
|
35 |
fixes x :: "'a::archimedean_field" shows "\<exists>n. x < of_nat n" |
|
36 |
proof - |
|
37 |
obtain z where "x < of_int z" using ex_less_of_int .. |
|
38 |
also have "\<dots> \<le> of_int (int (nat z))" by simp |
|
39 |
also have "\<dots> = of_nat (nat z)" by (simp only: of_int_of_nat_eq) |
|
40 |
finally show ?thesis .. |
|
41 |
qed |
|
42 |
||
43 |
lemma ex_le_of_nat: |
|
44 |
fixes x :: "'a::archimedean_field" shows "\<exists>n. x \<le> of_nat n" |
|
45 |
proof - |
|
46 |
obtain n where "x < of_nat n" using ex_less_of_nat .. |
|
47 |
then have "x \<le> of_nat n" by simp |
|
48 |
then show ?thesis .. |
|
49 |
qed |
|
50 |
||
51 |
text {* Archimedean fields have no infinitesimal elements. *} |
|
52 |
||
53 |
lemma ex_inverse_of_nat_Suc_less: |
|
54 |
fixes x :: "'a::archimedean_field" |
|
55 |
assumes "0 < x" shows "\<exists>n. inverse (of_nat (Suc n)) < x" |
|
56 |
proof - |
|
57 |
from `0 < x` have "0 < inverse x" |
|
58 |
by (rule positive_imp_inverse_positive) |
|
59 |
obtain n where "inverse x < of_nat n" |
|
60 |
using ex_less_of_nat .. |
|
61 |
then obtain m where "inverse x < of_nat (Suc m)" |
|
62 |
using `0 < inverse x` by (cases n) (simp_all del: of_nat_Suc) |
|
63 |
then have "inverse (of_nat (Suc m)) < inverse (inverse x)" |
|
64 |
using `0 < inverse x` by (rule less_imp_inverse_less) |
|
65 |
then have "inverse (of_nat (Suc m)) < x" |
|
66 |
using `0 < x` by (simp add: nonzero_inverse_inverse_eq) |
|
67 |
then show ?thesis .. |
|
68 |
qed |
|
69 |
||
70 |
lemma ex_inverse_of_nat_less: |
|
71 |
fixes x :: "'a::archimedean_field" |
|
72 |
assumes "0 < x" shows "\<exists>n>0. inverse (of_nat n) < x" |
|
73 |
using ex_inverse_of_nat_Suc_less [OF `0 < x`] by auto |
|
74 |
||
75 |
lemma ex_less_of_nat_mult: |
|
76 |
fixes x :: "'a::archimedean_field" |
|
77 |
assumes "0 < x" shows "\<exists>n. y < of_nat n * x" |
|
78 |
proof - |
|
79 |
obtain n where "y / x < of_nat n" using ex_less_of_nat .. |
|
80 |
with `0 < x` have "y < of_nat n * x" by (simp add: pos_divide_less_eq) |
|
81 |
then show ?thesis .. |
|
82 |
qed |
|
83 |
||
84 |
||
85 |
subsection {* Existence and uniqueness of floor function *} |
|
86 |
||
87 |
lemma exists_least_lemma: |
|
88 |
assumes "\<not> P 0" and "\<exists>n. P n" |
|
89 |
shows "\<exists>n. \<not> P n \<and> P (Suc n)" |
|
90 |
proof - |
|
91 |
from `\<exists>n. P n` have "P (Least P)" by (rule LeastI_ex) |
|
92 |
with `\<not> P 0` obtain n where "Least P = Suc n" |
|
93 |
by (cases "Least P") auto |
|
94 |
then have "n < Least P" by simp |
|
95 |
then have "\<not> P n" by (rule not_less_Least) |
|
96 |
then have "\<not> P n \<and> P (Suc n)" |
|
97 |
using `P (Least P)` `Least P = Suc n` by simp |
|
98 |
then show ?thesis .. |
|
99 |
qed |
|
100 |
||
101 |
lemma floor_exists: |
|
102 |
fixes x :: "'a::archimedean_field" |
|
103 |
shows "\<exists>z. of_int z \<le> x \<and> x < of_int (z + 1)" |
|
104 |
proof (cases) |
|
105 |
assume "0 \<le> x" |
|
106 |
then have "\<not> x < of_nat 0" by simp |
|
107 |
then have "\<exists>n. \<not> x < of_nat n \<and> x < of_nat (Suc n)" |
|
108 |
using ex_less_of_nat by (rule exists_least_lemma) |
|
109 |
then obtain n where "\<not> x < of_nat n \<and> x < of_nat (Suc n)" .. |
|
110 |
then have "of_int (int n) \<le> x \<and> x < of_int (int n + 1)" by simp |
|
111 |
then show ?thesis .. |
|
112 |
next |
|
113 |
assume "\<not> 0 \<le> x" |
|
114 |
then have "\<not> - x \<le> of_nat 0" by simp |
|
115 |
then have "\<exists>n. \<not> - x \<le> of_nat n \<and> - x \<le> of_nat (Suc n)" |
|
116 |
using ex_le_of_nat by (rule exists_least_lemma) |
|
117 |
then obtain n where "\<not> - x \<le> of_nat n \<and> - x \<le> of_nat (Suc n)" .. |
|
118 |
then have "of_int (- int n - 1) \<le> x \<and> x < of_int (- int n - 1 + 1)" by simp |
|
119 |
then show ?thesis .. |
|
120 |
qed |
|
121 |
||
122 |
lemma floor_exists1: |
|
123 |
fixes x :: "'a::archimedean_field" |
|
124 |
shows "\<exists>!z. of_int z \<le> x \<and> x < of_int (z + 1)" |
|
125 |
proof (rule ex_ex1I) |
|
126 |
show "\<exists>z. of_int z \<le> x \<and> x < of_int (z + 1)" |
|
127 |
by (rule floor_exists) |
|
128 |
next |
|
129 |
fix y z assume |
|
130 |
"of_int y \<le> x \<and> x < of_int (y + 1)" |
|
131 |
"of_int z \<le> x \<and> x < of_int (z + 1)" |
|
54281 | 132 |
with le_less_trans [of "of_int y" "x" "of_int (z + 1)"] |
133 |
le_less_trans [of "of_int z" "x" "of_int (y + 1)"] |
|
30096 | 134 |
show "y = z" by (simp del: of_int_add) |
135 |
qed |
|
136 |
||
137 |
||
138 |
subsection {* Floor function *} |
|
139 |
||
43732
6b2bdc57155b
adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents:
43704
diff
changeset
|
140 |
class floor_ceiling = archimedean_field + |
6b2bdc57155b
adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents:
43704
diff
changeset
|
141 |
fixes floor :: "'a \<Rightarrow> int" |
6b2bdc57155b
adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents:
43704
diff
changeset
|
142 |
assumes floor_correct: "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)" |
30096 | 143 |
|
144 |
notation (xsymbols) |
|
145 |
floor ("\<lfloor>_\<rfloor>") |
|
146 |
||
147 |
notation (HTML output) |
|
148 |
floor ("\<lfloor>_\<rfloor>") |
|
149 |
||
150 |
lemma floor_unique: "\<lbrakk>of_int z \<le> x; x < of_int z + 1\<rbrakk> \<Longrightarrow> floor x = z" |
|
151 |
using floor_correct [of x] floor_exists1 [of x] by auto |
|
152 |
||
153 |
lemma of_int_floor_le: "of_int (floor x) \<le> x" |
|
154 |
using floor_correct .. |
|
155 |
||
156 |
lemma le_floor_iff: "z \<le> floor x \<longleftrightarrow> of_int z \<le> x" |
|
157 |
proof |
|
158 |
assume "z \<le> floor x" |
|
159 |
then have "(of_int z :: 'a) \<le> of_int (floor x)" by simp |
|
160 |
also have "of_int (floor x) \<le> x" by (rule of_int_floor_le) |
|
161 |
finally show "of_int z \<le> x" . |
|
162 |
next |
|
163 |
assume "of_int z \<le> x" |
|
164 |
also have "x < of_int (floor x + 1)" using floor_correct .. |
|
165 |
finally show "z \<le> floor x" by (simp del: of_int_add) |
|
166 |
qed |
|
167 |
||
168 |
lemma floor_less_iff: "floor x < z \<longleftrightarrow> x < of_int z" |
|
169 |
by (simp add: not_le [symmetric] le_floor_iff) |
|
170 |
||
171 |
lemma less_floor_iff: "z < floor x \<longleftrightarrow> of_int z + 1 \<le> x" |
|
172 |
using le_floor_iff [of "z + 1" x] by auto |
|
173 |
||
174 |
lemma floor_le_iff: "floor x \<le> z \<longleftrightarrow> x < of_int z + 1" |
|
175 |
by (simp add: not_less [symmetric] less_floor_iff) |
|
176 |
||
177 |
lemma floor_mono: assumes "x \<le> y" shows "floor x \<le> floor y" |
|
178 |
proof - |
|
179 |
have "of_int (floor x) \<le> x" by (rule of_int_floor_le) |
|
180 |
also note `x \<le> y` |
|
181 |
finally show ?thesis by (simp add: le_floor_iff) |
|
182 |
qed |
|
183 |
||
184 |
lemma floor_less_cancel: "floor x < floor y \<Longrightarrow> x < y" |
|
185 |
by (auto simp add: not_le [symmetric] floor_mono) |
|
186 |
||
187 |
lemma floor_of_int [simp]: "floor (of_int z) = z" |
|
188 |
by (rule floor_unique) simp_all |
|
189 |
||
190 |
lemma floor_of_nat [simp]: "floor (of_nat n) = int n" |
|
191 |
using floor_of_int [of "of_nat n"] by simp |
|
192 |
||
47307
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
huffman
parents:
47108
diff
changeset
|
193 |
lemma le_floor_add: "floor x + floor y \<le> floor (x + y)" |
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
huffman
parents:
47108
diff
changeset
|
194 |
by (simp only: le_floor_iff of_int_add add_mono of_int_floor_le) |
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
huffman
parents:
47108
diff
changeset
|
195 |
|
30096 | 196 |
text {* Floor with numerals *} |
197 |
||
198 |
lemma floor_zero [simp]: "floor 0 = 0" |
|
199 |
using floor_of_int [of 0] by simp |
|
200 |
||
201 |
lemma floor_one [simp]: "floor 1 = 1" |
|
202 |
using floor_of_int [of 1] by simp |
|
203 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
204 |
lemma floor_numeral [simp]: "floor (numeral v) = numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
205 |
using floor_of_int [of "numeral v"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
206 |
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
207 |
lemma floor_neg_numeral [simp]: "floor (- numeral v) = - numeral v" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
208 |
using floor_of_int [of "- numeral v"] by simp |
30096 | 209 |
|
210 |
lemma zero_le_floor [simp]: "0 \<le> floor x \<longleftrightarrow> 0 \<le> x" |
|
211 |
by (simp add: le_floor_iff) |
|
212 |
||
213 |
lemma one_le_floor [simp]: "1 \<le> floor x \<longleftrightarrow> 1 \<le> x" |
|
214 |
by (simp add: le_floor_iff) |
|
215 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
216 |
lemma numeral_le_floor [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
217 |
"numeral v \<le> floor x \<longleftrightarrow> numeral v \<le> x" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
218 |
by (simp add: le_floor_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
219 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
220 |
lemma neg_numeral_le_floor [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
221 |
"- numeral v \<le> floor x \<longleftrightarrow> - numeral v \<le> x" |
30096 | 222 |
by (simp add: le_floor_iff) |
223 |
||
224 |
lemma zero_less_floor [simp]: "0 < floor x \<longleftrightarrow> 1 \<le> x" |
|
225 |
by (simp add: less_floor_iff) |
|
226 |
||
227 |
lemma one_less_floor [simp]: "1 < floor x \<longleftrightarrow> 2 \<le> x" |
|
228 |
by (simp add: less_floor_iff) |
|
229 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
230 |
lemma numeral_less_floor [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
231 |
"numeral v < floor x \<longleftrightarrow> numeral v + 1 \<le> x" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
232 |
by (simp add: less_floor_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
233 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
234 |
lemma neg_numeral_less_floor [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
235 |
"- numeral v < floor x \<longleftrightarrow> - numeral v + 1 \<le> x" |
30096 | 236 |
by (simp add: less_floor_iff) |
237 |
||
238 |
lemma floor_le_zero [simp]: "floor x \<le> 0 \<longleftrightarrow> x < 1" |
|
239 |
by (simp add: floor_le_iff) |
|
240 |
||
241 |
lemma floor_le_one [simp]: "floor x \<le> 1 \<longleftrightarrow> x < 2" |
|
242 |
by (simp add: floor_le_iff) |
|
243 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
244 |
lemma floor_le_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
245 |
"floor x \<le> numeral v \<longleftrightarrow> x < numeral v + 1" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
246 |
by (simp add: floor_le_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
247 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
248 |
lemma floor_le_neg_numeral [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
249 |
"floor x \<le> - numeral v \<longleftrightarrow> x < - numeral v + 1" |
30096 | 250 |
by (simp add: floor_le_iff) |
251 |
||
252 |
lemma floor_less_zero [simp]: "floor x < 0 \<longleftrightarrow> x < 0" |
|
253 |
by (simp add: floor_less_iff) |
|
254 |
||
255 |
lemma floor_less_one [simp]: "floor x < 1 \<longleftrightarrow> x < 1" |
|
256 |
by (simp add: floor_less_iff) |
|
257 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
258 |
lemma floor_less_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
259 |
"floor x < numeral v \<longleftrightarrow> x < numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
260 |
by (simp add: floor_less_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
261 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
262 |
lemma floor_less_neg_numeral [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
263 |
"floor x < - numeral v \<longleftrightarrow> x < - numeral v" |
30096 | 264 |
by (simp add: floor_less_iff) |
265 |
||
266 |
text {* Addition and subtraction of integers *} |
|
267 |
||
268 |
lemma floor_add_of_int [simp]: "floor (x + of_int z) = floor x + z" |
|
269 |
using floor_correct [of x] by (simp add: floor_unique) |
|
270 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
271 |
lemma floor_add_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
272 |
"floor (x + numeral v) = floor x + numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
273 |
using floor_add_of_int [of x "numeral v"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
274 |
|
30096 | 275 |
lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1" |
276 |
using floor_add_of_int [of x 1] by simp |
|
277 |
||
278 |
lemma floor_diff_of_int [simp]: "floor (x - of_int z) = floor x - z" |
|
279 |
using floor_add_of_int [of x "- z"] by (simp add: algebra_simps) |
|
280 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
281 |
lemma floor_diff_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
282 |
"floor (x - numeral v) = floor x - numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
283 |
using floor_diff_of_int [of x "numeral v"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
284 |
|
30096 | 285 |
lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1" |
286 |
using floor_diff_of_int [of x 1] by simp |
|
287 |
||
288 |
||
289 |
subsection {* Ceiling function *} |
|
290 |
||
291 |
definition |
|
43732
6b2bdc57155b
adding a floor_ceiling type class for different instantiations of floor (changeset from Brian Huffman)
bulwahn
parents:
43704
diff
changeset
|
292 |
ceiling :: "'a::floor_ceiling \<Rightarrow> int" where |
43733
a6ca7b83612f
adding code equations to execute floor and ceiling on rational and real numbers
bulwahn
parents:
43732
diff
changeset
|
293 |
"ceiling x = - floor (- x)" |
30096 | 294 |
|
295 |
notation (xsymbols) |
|
296 |
ceiling ("\<lceil>_\<rceil>") |
|
297 |
||
298 |
notation (HTML output) |
|
299 |
ceiling ("\<lceil>_\<rceil>") |
|
300 |
||
301 |
lemma ceiling_correct: "of_int (ceiling x) - 1 < x \<and> x \<le> of_int (ceiling x)" |
|
302 |
unfolding ceiling_def using floor_correct [of "- x"] by simp |
|
303 |
||
304 |
lemma ceiling_unique: "\<lbrakk>of_int z - 1 < x; x \<le> of_int z\<rbrakk> \<Longrightarrow> ceiling x = z" |
|
305 |
unfolding ceiling_def using floor_unique [of "- z" "- x"] by simp |
|
306 |
||
307 |
lemma le_of_int_ceiling: "x \<le> of_int (ceiling x)" |
|
308 |
using ceiling_correct .. |
|
309 |
||
310 |
lemma ceiling_le_iff: "ceiling x \<le> z \<longleftrightarrow> x \<le> of_int z" |
|
311 |
unfolding ceiling_def using le_floor_iff [of "- z" "- x"] by auto |
|
312 |
||
313 |
lemma less_ceiling_iff: "z < ceiling x \<longleftrightarrow> of_int z < x" |
|
314 |
by (simp add: not_le [symmetric] ceiling_le_iff) |
|
315 |
||
316 |
lemma ceiling_less_iff: "ceiling x < z \<longleftrightarrow> x \<le> of_int z - 1" |
|
317 |
using ceiling_le_iff [of x "z - 1"] by simp |
|
318 |
||
319 |
lemma le_ceiling_iff: "z \<le> ceiling x \<longleftrightarrow> of_int z - 1 < x" |
|
320 |
by (simp add: not_less [symmetric] ceiling_less_iff) |
|
321 |
||
322 |
lemma ceiling_mono: "x \<ge> y \<Longrightarrow> ceiling x \<ge> ceiling y" |
|
323 |
unfolding ceiling_def by (simp add: floor_mono) |
|
324 |
||
325 |
lemma ceiling_less_cancel: "ceiling x < ceiling y \<Longrightarrow> x < y" |
|
326 |
by (auto simp add: not_le [symmetric] ceiling_mono) |
|
327 |
||
328 |
lemma ceiling_of_int [simp]: "ceiling (of_int z) = z" |
|
329 |
by (rule ceiling_unique) simp_all |
|
330 |
||
331 |
lemma ceiling_of_nat [simp]: "ceiling (of_nat n) = int n" |
|
332 |
using ceiling_of_int [of "of_nat n"] by simp |
|
333 |
||
47307
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
huffman
parents:
47108
diff
changeset
|
334 |
lemma ceiling_add_le: "ceiling (x + y) \<le> ceiling x + ceiling y" |
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
huffman
parents:
47108
diff
changeset
|
335 |
by (simp only: ceiling_le_iff of_int_add add_mono le_of_int_ceiling) |
5e5ca36692b3
add floor/ceiling lemmas suggested by René Thiemann
huffman
parents:
47108
diff
changeset
|
336 |
|
30096 | 337 |
text {* Ceiling with numerals *} |
338 |
||
339 |
lemma ceiling_zero [simp]: "ceiling 0 = 0" |
|
340 |
using ceiling_of_int [of 0] by simp |
|
341 |
||
342 |
lemma ceiling_one [simp]: "ceiling 1 = 1" |
|
343 |
using ceiling_of_int [of 1] by simp |
|
344 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
345 |
lemma ceiling_numeral [simp]: "ceiling (numeral v) = numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
346 |
using ceiling_of_int [of "numeral v"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
347 |
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
348 |
lemma ceiling_neg_numeral [simp]: "ceiling (- numeral v) = - numeral v" |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
349 |
using ceiling_of_int [of "- numeral v"] by simp |
30096 | 350 |
|
351 |
lemma ceiling_le_zero [simp]: "ceiling x \<le> 0 \<longleftrightarrow> x \<le> 0" |
|
352 |
by (simp add: ceiling_le_iff) |
|
353 |
||
354 |
lemma ceiling_le_one [simp]: "ceiling x \<le> 1 \<longleftrightarrow> x \<le> 1" |
|
355 |
by (simp add: ceiling_le_iff) |
|
356 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
357 |
lemma ceiling_le_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
358 |
"ceiling x \<le> numeral v \<longleftrightarrow> x \<le> numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
359 |
by (simp add: ceiling_le_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
360 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
361 |
lemma ceiling_le_neg_numeral [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
362 |
"ceiling x \<le> - numeral v \<longleftrightarrow> x \<le> - numeral v" |
30096 | 363 |
by (simp add: ceiling_le_iff) |
364 |
||
365 |
lemma ceiling_less_zero [simp]: "ceiling x < 0 \<longleftrightarrow> x \<le> -1" |
|
366 |
by (simp add: ceiling_less_iff) |
|
367 |
||
368 |
lemma ceiling_less_one [simp]: "ceiling x < 1 \<longleftrightarrow> x \<le> 0" |
|
369 |
by (simp add: ceiling_less_iff) |
|
370 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
371 |
lemma ceiling_less_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
372 |
"ceiling x < numeral v \<longleftrightarrow> x \<le> numeral v - 1" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
373 |
by (simp add: ceiling_less_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
374 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
375 |
lemma ceiling_less_neg_numeral [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
376 |
"ceiling x < - numeral v \<longleftrightarrow> x \<le> - numeral v - 1" |
30096 | 377 |
by (simp add: ceiling_less_iff) |
378 |
||
379 |
lemma zero_le_ceiling [simp]: "0 \<le> ceiling x \<longleftrightarrow> -1 < x" |
|
380 |
by (simp add: le_ceiling_iff) |
|
381 |
||
382 |
lemma one_le_ceiling [simp]: "1 \<le> ceiling x \<longleftrightarrow> 0 < x" |
|
383 |
by (simp add: le_ceiling_iff) |
|
384 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
385 |
lemma numeral_le_ceiling [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
386 |
"numeral v \<le> ceiling x \<longleftrightarrow> numeral v - 1 < x" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
387 |
by (simp add: le_ceiling_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
388 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
389 |
lemma neg_numeral_le_ceiling [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
390 |
"- numeral v \<le> ceiling x \<longleftrightarrow> - numeral v - 1 < x" |
30096 | 391 |
by (simp add: le_ceiling_iff) |
392 |
||
393 |
lemma zero_less_ceiling [simp]: "0 < ceiling x \<longleftrightarrow> 0 < x" |
|
394 |
by (simp add: less_ceiling_iff) |
|
395 |
||
396 |
lemma one_less_ceiling [simp]: "1 < ceiling x \<longleftrightarrow> 1 < x" |
|
397 |
by (simp add: less_ceiling_iff) |
|
398 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
399 |
lemma numeral_less_ceiling [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
400 |
"numeral v < ceiling x \<longleftrightarrow> numeral v < x" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
401 |
by (simp add: less_ceiling_iff) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
402 |
|
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
403 |
lemma neg_numeral_less_ceiling [simp]: |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54281
diff
changeset
|
404 |
"- numeral v < ceiling x \<longleftrightarrow> - numeral v < x" |
30096 | 405 |
by (simp add: less_ceiling_iff) |
406 |
||
407 |
text {* Addition and subtraction of integers *} |
|
408 |
||
409 |
lemma ceiling_add_of_int [simp]: "ceiling (x + of_int z) = ceiling x + z" |
|
410 |
using ceiling_correct [of x] by (simp add: ceiling_unique) |
|
411 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
412 |
lemma ceiling_add_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
413 |
"ceiling (x + numeral v) = ceiling x + numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
414 |
using ceiling_add_of_int [of x "numeral v"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
415 |
|
30096 | 416 |
lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1" |
417 |
using ceiling_add_of_int [of x 1] by simp |
|
418 |
||
419 |
lemma ceiling_diff_of_int [simp]: "ceiling (x - of_int z) = ceiling x - z" |
|
420 |
using ceiling_add_of_int [of x "- z"] by (simp add: algebra_simps) |
|
421 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
422 |
lemma ceiling_diff_numeral [simp]: |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
423 |
"ceiling (x - numeral v) = ceiling x - numeral v" |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
424 |
using ceiling_diff_of_int [of x "numeral v"] by simp |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
43733
diff
changeset
|
425 |
|
30096 | 426 |
lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1" |
427 |
using ceiling_diff_of_int [of x 1] by simp |
|
428 |
||
47592 | 429 |
lemma ceiling_diff_floor_le_1: "ceiling x - floor x \<le> 1" |
430 |
proof - |
|
431 |
have "of_int \<lceil>x\<rceil> - 1 < x" |
|
432 |
using ceiling_correct[of x] by simp |
|
433 |
also have "x < of_int \<lfloor>x\<rfloor> + 1" |
|
434 |
using floor_correct[of x] by simp_all |
|
435 |
finally have "of_int (\<lceil>x\<rceil> - \<lfloor>x\<rfloor>) < (of_int 2::'a)" |
|
436 |
by simp |
|
437 |
then show ?thesis |
|
438 |
unfolding of_int_less_iff by simp |
|
439 |
qed |
|
30096 | 440 |
|
441 |
subsection {* Negation *} |
|
442 |
||
30102 | 443 |
lemma floor_minus: "floor (- x) = - ceiling x" |
30096 | 444 |
unfolding ceiling_def by simp |
445 |
||
30102 | 446 |
lemma ceiling_minus: "ceiling (- x) = - floor x" |
30096 | 447 |
unfolding ceiling_def by simp |
448 |
||
449 |
end |