1 (* Title: HOL/Archimedean_Field.thy
5 header {* Archimedean Fields, Floor and Ceiling Functions *}
7 theory Archimedean_Field
11 subsection {* Class of Archimedean fields *}
13 text {* Archimedean fields have no infinite elements. *}
15 class archimedean_field = linordered_field +
16 assumes ex_le_of_int: "\<exists>z. x \<le> of_int z"
19 fixes x :: "'a::archimedean_field" shows "\<exists>z. x < of_int z"
21 from ex_le_of_int obtain z where "x \<le> of_int z" ..
22 then have "x < of_int (z + 1)" by simp
27 fixes x :: "'a::archimedean_field" shows "\<exists>z. of_int z < x"
29 from ex_less_of_int obtain z where "- x < of_int z" ..
30 then have "of_int (- z) < x" by simp
35 fixes x :: "'a::archimedean_field" shows "\<exists>n. x < of_nat n"
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 ..
44 fixes x :: "'a::archimedean_field" shows "\<exists>n. x \<le> of_nat n"
46 obtain n where "x < of_nat n" using ex_less_of_nat ..
47 then have "x \<le> of_nat n" by simp
51 text {* Archimedean fields have no infinitesimal elements. *}
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"
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)
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
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"
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)
85 subsection {* Existence and uniqueness of floor function *}
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)"
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
102 fixes x :: "'a::archimedean_field"
103 shows "\<exists>z. of_int z \<le> x \<and> x < of_int (z + 1)"
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
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
123 fixes x :: "'a::archimedean_field"
124 shows "\<exists>!z. of_int z \<le> x \<and> x < of_int (z + 1)"
126 show "\<exists>z. of_int z \<le> x \<and> x < of_int (z + 1)"
127 by (rule floor_exists)
130 "of_int y \<le> x \<and> x < of_int (y + 1)"
131 "of_int z \<le> x \<and> x < of_int (z + 1)"
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)"]
134 show "y = z" by (simp del: of_int_add)
138 subsection {* Floor function *}
140 class floor_ceiling = archimedean_field +
141 fixes floor :: "'a \<Rightarrow> int"
142 assumes floor_correct: "of_int (floor x) \<le> x \<and> x < of_int (floor x + 1)"
145 floor ("\<lfloor>_\<rfloor>")
147 notation (HTML output)
148 floor ("\<lfloor>_\<rfloor>")
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
153 lemma of_int_floor_le: "of_int (floor x) \<le> x"
154 using floor_correct ..
156 lemma le_floor_iff: "z \<le> floor x \<longleftrightarrow> of_int z \<le> x"
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" .
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)
168 lemma floor_less_iff: "floor x < z \<longleftrightarrow> x < of_int z"
169 by (simp add: not_le [symmetric] le_floor_iff)
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
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)
177 lemma floor_mono: assumes "x \<le> y" shows "floor x \<le> floor y"
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)
184 lemma floor_less_cancel: "floor x < floor y \<Longrightarrow> x < y"
185 by (auto simp add: not_le [symmetric] floor_mono)
187 lemma floor_of_int [simp]: "floor (of_int z) = z"
188 by (rule floor_unique) simp_all
190 lemma floor_of_nat [simp]: "floor (of_nat n) = int n"
191 using floor_of_int [of "of_nat n"] by simp
193 lemma le_floor_add: "floor x + floor y \<le> floor (x + y)"
194 by (simp only: le_floor_iff of_int_add add_mono of_int_floor_le)
196 text {* Floor with numerals *}
198 lemma floor_zero [simp]: "floor 0 = 0"
199 using floor_of_int [of 0] by simp
201 lemma floor_one [simp]: "floor 1 = 1"
202 using floor_of_int [of 1] by simp
204 lemma floor_numeral [simp]: "floor (numeral v) = numeral v"
205 using floor_of_int [of "numeral v"] by simp
207 lemma floor_neg_numeral [simp]: "floor (- numeral v) = - numeral v"
208 using floor_of_int [of "- numeral v"] by simp
210 lemma zero_le_floor [simp]: "0 \<le> floor x \<longleftrightarrow> 0 \<le> x"
211 by (simp add: le_floor_iff)
213 lemma one_le_floor [simp]: "1 \<le> floor x \<longleftrightarrow> 1 \<le> x"
214 by (simp add: le_floor_iff)
216 lemma numeral_le_floor [simp]:
217 "numeral v \<le> floor x \<longleftrightarrow> numeral v \<le> x"
218 by (simp add: le_floor_iff)
220 lemma neg_numeral_le_floor [simp]:
221 "- numeral v \<le> floor x \<longleftrightarrow> - numeral v \<le> x"
222 by (simp add: le_floor_iff)
224 lemma zero_less_floor [simp]: "0 < floor x \<longleftrightarrow> 1 \<le> x"
225 by (simp add: less_floor_iff)
227 lemma one_less_floor [simp]: "1 < floor x \<longleftrightarrow> 2 \<le> x"
228 by (simp add: less_floor_iff)
230 lemma numeral_less_floor [simp]:
231 "numeral v < floor x \<longleftrightarrow> numeral v + 1 \<le> x"
232 by (simp add: less_floor_iff)
234 lemma neg_numeral_less_floor [simp]:
235 "- numeral v < floor x \<longleftrightarrow> - numeral v + 1 \<le> x"
236 by (simp add: less_floor_iff)
238 lemma floor_le_zero [simp]: "floor x \<le> 0 \<longleftrightarrow> x < 1"
239 by (simp add: floor_le_iff)
241 lemma floor_le_one [simp]: "floor x \<le> 1 \<longleftrightarrow> x < 2"
242 by (simp add: floor_le_iff)
244 lemma floor_le_numeral [simp]:
245 "floor x \<le> numeral v \<longleftrightarrow> x < numeral v + 1"
246 by (simp add: floor_le_iff)
248 lemma floor_le_neg_numeral [simp]:
249 "floor x \<le> - numeral v \<longleftrightarrow> x < - numeral v + 1"
250 by (simp add: floor_le_iff)
252 lemma floor_less_zero [simp]: "floor x < 0 \<longleftrightarrow> x < 0"
253 by (simp add: floor_less_iff)
255 lemma floor_less_one [simp]: "floor x < 1 \<longleftrightarrow> x < 1"
256 by (simp add: floor_less_iff)
258 lemma floor_less_numeral [simp]:
259 "floor x < numeral v \<longleftrightarrow> x < numeral v"
260 by (simp add: floor_less_iff)
262 lemma floor_less_neg_numeral [simp]:
263 "floor x < - numeral v \<longleftrightarrow> x < - numeral v"
264 by (simp add: floor_less_iff)
266 text {* Addition and subtraction of integers *}
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)
271 lemma floor_add_numeral [simp]:
272 "floor (x + numeral v) = floor x + numeral v"
273 using floor_add_of_int [of x "numeral v"] by simp
275 lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1"
276 using floor_add_of_int [of x 1] by simp
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)
281 lemma floor_diff_numeral [simp]:
282 "floor (x - numeral v) = floor x - numeral v"
283 using floor_diff_of_int [of x "numeral v"] by simp
285 lemma floor_diff_one [simp]: "floor (x - 1) = floor x - 1"
286 using floor_diff_of_int [of x 1] by simp
289 subsection {* Ceiling function *}
292 ceiling :: "'a::floor_ceiling \<Rightarrow> int" where
293 "ceiling x = - floor (- x)"
296 ceiling ("\<lceil>_\<rceil>")
298 notation (HTML output)
299 ceiling ("\<lceil>_\<rceil>")
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
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
307 lemma le_of_int_ceiling: "x \<le> of_int (ceiling x)"
308 using ceiling_correct ..
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
313 lemma less_ceiling_iff: "z < ceiling x \<longleftrightarrow> of_int z < x"
314 by (simp add: not_le [symmetric] ceiling_le_iff)
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
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)
322 lemma ceiling_mono: "x \<ge> y \<Longrightarrow> ceiling x \<ge> ceiling y"
323 unfolding ceiling_def by (simp add: floor_mono)
325 lemma ceiling_less_cancel: "ceiling x < ceiling y \<Longrightarrow> x < y"
326 by (auto simp add: not_le [symmetric] ceiling_mono)
328 lemma ceiling_of_int [simp]: "ceiling (of_int z) = z"
329 by (rule ceiling_unique) simp_all
331 lemma ceiling_of_nat [simp]: "ceiling (of_nat n) = int n"
332 using ceiling_of_int [of "of_nat n"] by simp
334 lemma ceiling_add_le: "ceiling (x + y) \<le> ceiling x + ceiling y"
335 by (simp only: ceiling_le_iff of_int_add add_mono le_of_int_ceiling)
337 text {* Ceiling with numerals *}
339 lemma ceiling_zero [simp]: "ceiling 0 = 0"
340 using ceiling_of_int [of 0] by simp
342 lemma ceiling_one [simp]: "ceiling 1 = 1"
343 using ceiling_of_int [of 1] by simp
345 lemma ceiling_numeral [simp]: "ceiling (numeral v) = numeral v"
346 using ceiling_of_int [of "numeral v"] by simp
348 lemma ceiling_neg_numeral [simp]: "ceiling (- numeral v) = - numeral v"
349 using ceiling_of_int [of "- numeral v"] by simp
351 lemma ceiling_le_zero [simp]: "ceiling x \<le> 0 \<longleftrightarrow> x \<le> 0"
352 by (simp add: ceiling_le_iff)
354 lemma ceiling_le_one [simp]: "ceiling x \<le> 1 \<longleftrightarrow> x \<le> 1"
355 by (simp add: ceiling_le_iff)
357 lemma ceiling_le_numeral [simp]:
358 "ceiling x \<le> numeral v \<longleftrightarrow> x \<le> numeral v"
359 by (simp add: ceiling_le_iff)
361 lemma ceiling_le_neg_numeral [simp]:
362 "ceiling x \<le> - numeral v \<longleftrightarrow> x \<le> - numeral v"
363 by (simp add: ceiling_le_iff)
365 lemma ceiling_less_zero [simp]: "ceiling x < 0 \<longleftrightarrow> x \<le> -1"
366 by (simp add: ceiling_less_iff)
368 lemma ceiling_less_one [simp]: "ceiling x < 1 \<longleftrightarrow> x \<le> 0"
369 by (simp add: ceiling_less_iff)
371 lemma ceiling_less_numeral [simp]:
372 "ceiling x < numeral v \<longleftrightarrow> x \<le> numeral v - 1"
373 by (simp add: ceiling_less_iff)
375 lemma ceiling_less_neg_numeral [simp]:
376 "ceiling x < - numeral v \<longleftrightarrow> x \<le> - numeral v - 1"
377 by (simp add: ceiling_less_iff)
379 lemma zero_le_ceiling [simp]: "0 \<le> ceiling x \<longleftrightarrow> -1 < x"
380 by (simp add: le_ceiling_iff)
382 lemma one_le_ceiling [simp]: "1 \<le> ceiling x \<longleftrightarrow> 0 < x"
383 by (simp add: le_ceiling_iff)
385 lemma numeral_le_ceiling [simp]:
386 "numeral v \<le> ceiling x \<longleftrightarrow> numeral v - 1 < x"
387 by (simp add: le_ceiling_iff)
389 lemma neg_numeral_le_ceiling [simp]:
390 "- numeral v \<le> ceiling x \<longleftrightarrow> - numeral v - 1 < x"
391 by (simp add: le_ceiling_iff)
393 lemma zero_less_ceiling [simp]: "0 < ceiling x \<longleftrightarrow> 0 < x"
394 by (simp add: less_ceiling_iff)
396 lemma one_less_ceiling [simp]: "1 < ceiling x \<longleftrightarrow> 1 < x"
397 by (simp add: less_ceiling_iff)
399 lemma numeral_less_ceiling [simp]:
400 "numeral v < ceiling x \<longleftrightarrow> numeral v < x"
401 by (simp add: less_ceiling_iff)
403 lemma neg_numeral_less_ceiling [simp]:
404 "- numeral v < ceiling x \<longleftrightarrow> - numeral v < x"
405 by (simp add: less_ceiling_iff)
407 text {* Addition and subtraction of integers *}
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)
412 lemma ceiling_add_numeral [simp]:
413 "ceiling (x + numeral v) = ceiling x + numeral v"
414 using ceiling_add_of_int [of x "numeral v"] by simp
416 lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1"
417 using ceiling_add_of_int [of x 1] by simp
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)
422 lemma ceiling_diff_numeral [simp]:
423 "ceiling (x - numeral v) = ceiling x - numeral v"
424 using ceiling_diff_of_int [of x "numeral v"] by simp
426 lemma ceiling_diff_one [simp]: "ceiling (x - 1) = ceiling x - 1"
427 using ceiling_diff_of_int [of x 1] by simp
429 lemma ceiling_diff_floor_le_1: "ceiling x - floor x \<le> 1"
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)"
438 unfolding of_int_less_iff by simp
441 subsection {* Negation *}
443 lemma floor_minus: "floor (- x) = - ceiling x"
444 unfolding ceiling_def by simp
446 lemma ceiling_minus: "ceiling (- x) = - floor x"
447 unfolding ceiling_def by simp