made repository layout more coherent with logical distribution structure; stripped some $Id$s
1 (* Title : HOL/RComplete.thy
2 Author : Jacques D. Fleuriot, University of Edinburgh
3 Author : Larry Paulson, University of Cambridge
4 Author : Jeremy Avigad, Carnegie Mellon University
5 Author : Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen
8 header {* Completeness of the Reals; Floor and Ceiling Functions *}
14 lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
18 subsection {* Completeness of Positive Reals *}
21 Supremum property for the set of positive reals
23 Let @{text "P"} be a non-empty set of positive reals, with an upper
24 bound @{text "y"}. Then @{text "P"} has a least upper bound
25 (written @{text "S"}).
27 FIXME: Can the premise be weakened to @{text "\<forall>x \<in> P. x\<le> y"}?
30 lemma posreal_complete:
31 assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
32 and not_empty_P: "\<exists>x. x \<in> P"
33 and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
34 shows "\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S)"
35 proof (rule exI, rule allI)
37 let ?pP = "{w. real_of_preal w \<in> P}"
39 show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"
41 assume neg_y: "\<not> 0 < y"
44 assume "\<exists>x\<in>P. y < x"
45 have "\<forall>x. y < real_of_preal x"
46 using neg_y by (rule real_less_all_real2)
47 thus "y < real_of_preal (psup ?pP)" ..
49 assume "y < real_of_preal (psup ?pP)"
50 obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
51 hence "0 < x" using positive_P by simp
52 hence "y < x" using neg_y by simp
53 thus "\<exists>x \<in> P. y < x" using x_in_P ..
58 then obtain py where y_is_py: "y = real_of_preal py"
59 by (auto simp add: real_gt_zero_preal_Ex)
61 obtain a where "a \<in> P" using not_empty_P ..
62 with positive_P have a_pos: "0 < a" ..
63 then obtain pa where "a = real_of_preal pa"
64 by (auto simp add: real_gt_zero_preal_Ex)
65 hence "pa \<in> ?pP" using `a \<in> P` by auto
66 hence pP_not_empty: "?pP \<noteq> {}" by auto
68 obtain sup where sup: "\<forall>x \<in> P. x < sup"
69 using upper_bound_Ex ..
70 from this and `a \<in> P` have "a < sup" ..
71 hence "0 < sup" using a_pos by arith
72 then obtain possup where "sup = real_of_preal possup"
73 by (auto simp add: real_gt_zero_preal_Ex)
74 hence "\<forall>X \<in> ?pP. X \<le> possup"
75 using sup by (auto simp add: real_of_preal_lessI)
76 with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"
77 by (rule preal_complete)
81 assume "\<exists>x \<in> P. y < x"
82 then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..
83 hence "0 < x" using pos_y by arith
84 then obtain px where x_is_px: "x = real_of_preal px"
85 by (auto simp add: real_gt_zero_preal_Ex)
87 have py_less_X: "\<exists>X \<in> ?pP. py < X"
89 show "py < px" using y_is_py and x_is_px and y_less_x
90 by (simp add: real_of_preal_lessI)
91 show "px \<in> ?pP" using x_in_P and x_is_px by simp
94 have "(\<exists>X \<in> ?pP. py < X) ==> (py < psup ?pP)"
96 hence "py < psup ?pP" using py_less_X by simp
97 thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
98 using y_is_py and pos_y by (simp add: real_of_preal_lessI)
100 assume y_less_psup: "y < real_of_preal (psup ?pP)"
102 hence "py < psup ?pP" using y_is_py
103 by (simp add: real_of_preal_lessI)
104 then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"
106 then obtain x where x_is_X: "x = real_of_preal X"
107 by (simp add: real_gt_zero_preal_Ex)
108 hence "y < x" using py_less_X and y_is_py
109 by (simp add: real_of_preal_lessI)
111 moreover have "x \<in> P" using x_is_X and X_in_pP by simp
113 ultimately show "\<exists> x \<in> P. y < x" ..
119 \medskip Completeness properties using @{text "isUb"}, @{text "isLub"} etc.
122 lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
123 apply (frule isLub_isUb)
124 apply (frule_tac x = y in isLub_isUb)
125 apply (blast intro!: order_antisym dest!: isLub_le_isUb)
130 \medskip Completeness theorem for the positive reals (again).
133 lemma posreals_complete:
134 assumes positive_S: "\<forall>x \<in> S. 0 < x"
135 and not_empty_S: "\<exists>x. x \<in> S"
136 and upper_bound_Ex: "\<exists>u. isUb (UNIV::real set) S u"
137 shows "\<exists>t. isLub (UNIV::real set) S t"
139 let ?pS = "{w. real_of_preal w \<in> S}"
141 obtain u where "isUb UNIV S u" using upper_bound_Ex ..
142 hence sup: "\<forall>x \<in> S. x \<le> u" by (simp add: isUb_def setle_def)
144 obtain x where x_in_S: "x \<in> S" using not_empty_S ..
145 hence x_gt_zero: "0 < x" using positive_S by simp
146 have "x \<le> u" using sup and x_in_S ..
147 hence "0 < u" using x_gt_zero by arith
149 then obtain pu where u_is_pu: "u = real_of_preal pu"
150 by (auto simp add: real_gt_zero_preal_Ex)
152 have pS_less_pu: "\<forall>pa \<in> ?pS. pa \<le> pu"
155 assume "pa \<in> ?pS"
156 then obtain a where "a \<in> S" and "a = real_of_preal pa"
158 moreover hence "a \<le> u" using sup by simp
159 ultimately show "pa \<le> pu"
160 using sup and u_is_pu by (simp add: real_of_preal_le_iff)
163 have "\<forall>y \<in> S. y \<le> real_of_preal (psup ?pS)"
166 assume y_in_S: "y \<in> S"
167 hence "0 < y" using positive_S by simp
168 then obtain py where y_is_py: "y = real_of_preal py"
169 by (auto simp add: real_gt_zero_preal_Ex)
170 hence py_in_pS: "py \<in> ?pS" using y_in_S by simp
171 with pS_less_pu have "py \<le> psup ?pS"
172 by (rule preal_psup_le)
173 thus "y \<le> real_of_preal (psup ?pS)"
174 using y_is_py by (simp add: real_of_preal_le_iff)
179 assume x_ub_S: "\<forall>y\<in>S. y \<le> x"
180 have "real_of_preal (psup ?pS) \<le> x"
182 obtain "s" where s_in_S: "s \<in> S" using not_empty_S ..
183 hence s_pos: "0 < s" using positive_S by simp
185 hence "\<exists> ps. s = real_of_preal ps" by (simp add: real_gt_zero_preal_Ex)
186 then obtain "ps" where s_is_ps: "s = real_of_preal ps" ..
187 hence ps_in_pS: "ps \<in> {w. real_of_preal w \<in> S}" using s_in_S by simp
189 from x_ub_S have "s \<le> x" using s_in_S ..
190 hence "0 < x" using s_pos by simp
191 hence "\<exists> px. x = real_of_preal px" by (simp add: real_gt_zero_preal_Ex)
192 then obtain "px" where x_is_px: "x = real_of_preal px" ..
194 have "\<forall>pe \<in> ?pS. pe \<le> px"
197 assume "pe \<in> ?pS"
198 hence "real_of_preal pe \<in> S" by simp
199 hence "real_of_preal pe \<le> x" using x_ub_S by simp
200 thus "pe \<le> px" using x_is_px by (simp add: real_of_preal_le_iff)
203 moreover have "?pS \<noteq> {}" using ps_in_pS by auto
204 ultimately have "(psup ?pS) \<le> px" by (simp add: psup_le_ub)
205 thus "real_of_preal (psup ?pS) \<le> x" using x_is_px by (simp add: real_of_preal_le_iff)
208 ultimately show "isLub UNIV S (real_of_preal (psup ?pS))"
209 by (simp add: isLub_def leastP_def isUb_def setle_def setge_def)
213 \medskip reals Completeness (again!)
216 lemma reals_complete:
217 assumes notempty_S: "\<exists>X. X \<in> S"
218 and exists_Ub: "\<exists>Y. isUb (UNIV::real set) S Y"
219 shows "\<exists>t. isLub (UNIV :: real set) S t"
221 obtain X where X_in_S: "X \<in> S" using notempty_S ..
222 obtain Y where Y_isUb: "isUb (UNIV::real set) S Y"
224 let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
228 assume "isUb (UNIV::real set) S x"
229 hence S_le_x: "\<forall> y \<in> S. y <= x"
230 by (simp add: isUb_def setle_def)
233 assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
234 hence "\<exists> x \<in> S. s = x + -X + 1" ..
235 then obtain x1 where "x1 \<in> S" and "s = x1 + (-X) + 1" ..
236 moreover hence "x1 \<le> x" using S_le_x by simp
237 ultimately have "s \<le> x + - X + 1" by arith
239 then have "isUb (UNIV::real set) ?SHIFT (x + (-X) + 1)"
240 by (auto simp add: isUb_def setle_def)
241 } note S_Ub_is_SHIFT_Ub = this
243 hence "isUb UNIV ?SHIFT (Y + (-X) + 1)" using Y_isUb by simp
244 hence "\<exists>Z. isUb UNIV ?SHIFT Z" ..
245 moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
246 moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
247 using X_in_S and Y_isUb by auto
248 ultimately obtain t where t_is_Lub: "isLub UNIV ?SHIFT t"
249 using posreals_complete [of ?SHIFT] by blast
253 show "isLub UNIV S (t + X + (-1))"
257 assume "isUb (UNIV::real set) S x"
258 hence "isUb (UNIV::real set) (?SHIFT) (x + (-X) + 1)"
259 using S_Ub_is_SHIFT_Ub by simp
260 hence "t \<le> (x + (-X) + 1)"
261 using t_is_Lub by (simp add: isLub_le_isUb)
262 hence "t + X + -1 \<le> x" by arith
264 then show "(t + X + -1) <=* Collect (isUb UNIV S)"
265 by (simp add: setgeI)
267 show "isUb UNIV S (t + X + -1)"
271 assume y_in_S: "y \<in> S"
272 have "y \<le> t + X + -1"
274 obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
275 hence "\<exists> x \<in> S. u = x + - X + 1" by simp
276 then obtain "x" where x_and_u: "u = x + - X + 1" ..
277 have u_le_t: "u \<le> t" using u_in_shift and t_is_Lub by (simp add: isLubD2)
282 moreover have "x = u + X + - 1" using x_and_u by arith
283 moreover have "u + X + - 1 \<le> t + X + -1" using u_le_t by arith
284 ultimately show "y \<le> t + X + -1" by arith
286 assume "~(y \<le> x)"
287 hence x_less_y: "x < y" by arith
289 have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
290 hence "0 < x + (-X) + 1" by simp
291 hence "0 < y + (-X) + 1" using x_less_y by arith
292 hence "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
293 hence "y + (-X) + 1 \<le> t" using t_is_Lub by (simp add: isLubD2)
298 then show ?thesis by (simp add: isUb_def setle_def)
305 subsection {* The Archimedean Property of the Reals *}
307 theorem reals_Archimedean:
308 assumes x_pos: "0 < x"
309 shows "\<exists>n. inverse (real (Suc n)) < x"
311 assume contr: "\<not> ?thesis"
312 have "\<forall>n. x * real (Suc n) <= 1"
315 from contr have "x \<le> inverse (real (Suc n))"
316 by (simp add: linorder_not_less)
317 hence "x \<le> (1 / (real (Suc n)))"
318 by (simp add: inverse_eq_divide)
319 moreover have "0 \<le> real (Suc n)"
320 by (rule real_of_nat_ge_zero)
321 ultimately have "x * real (Suc n) \<le> (1 / real (Suc n)) * real (Suc n)"
322 by (rule mult_right_mono)
323 thus "x * real (Suc n) \<le> 1" by simp
325 hence "{z. \<exists>n. z = x * (real (Suc n))} *<= 1"
326 by (simp add: setle_def, safe, rule spec)
327 hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (real (Suc n))} 1"
329 hence "\<exists>Y. isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} Y" ..
330 moreover have "\<exists>X. X \<in> {z. \<exists>n. z = x* (real (Suc n))}" by auto
331 ultimately have "\<exists>t. isLub UNIV {z. \<exists>n. z = x * real (Suc n)} t"
332 by (simp add: reals_complete)
333 then obtain "t" where
334 t_is_Lub: "isLub UNIV {z. \<exists>n. z = x * real (Suc n)} t" ..
336 have "\<forall>n::nat. x * real n \<le> t + - x"
339 from t_is_Lub have "x * real (Suc n) \<le> t"
340 by (simp add: isLubD2)
341 hence "x * (real n) + x \<le> t"
342 by (simp add: right_distrib real_of_nat_Suc)
343 thus "x * (real n) \<le> t + - x" by arith
346 hence "\<forall>m. x * real (Suc m) \<le> t + - x" by simp
347 hence "{z. \<exists>n. z = x * (real (Suc n))} *<= (t + - x)"
348 by (auto simp add: setle_def)
349 hence "isUb (UNIV::real set) {z. \<exists>n. z = x * (real (Suc n))} (t + (-x))"
351 hence "t \<le> t + - x"
352 using t_is_Lub by (simp add: isLub_le_isUb)
353 thus False using x_pos by arith
357 There must be other proofs, e.g. @{text "Suc"} of the largest
358 integer in the cut representing @{text "x"}.
361 lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
364 hence "x < real (1::nat)" by simp
367 assume "\<not> x \<le> 0"
368 hence x_greater_zero: "0 < x" by simp
369 hence "0 < inverse x" by simp
370 then obtain n where "inverse (real (Suc n)) < inverse x"
371 using reals_Archimedean by blast
372 hence "inverse (real (Suc n)) * x < inverse x * x"
373 using x_greater_zero by (rule mult_strict_right_mono)
374 hence "inverse (real (Suc n)) * x < 1"
375 using x_greater_zero by simp
376 hence "real (Suc n) * (inverse (real (Suc n)) * x) < real (Suc n) * 1"
377 by (rule mult_strict_left_mono) simp
378 hence "x < real (Suc n)"
379 by (simp add: ring_simps)
380 thus "\<exists>(n::nat). x < real n" ..
383 lemma reals_Archimedean3:
384 assumes x_greater_zero: "0 < x"
385 shows "\<forall>(y::real). \<exists>(n::nat). y < real n * x"
388 have x_not_zero: "x \<noteq> 0" using x_greater_zero by simp
389 obtain n where "y * inverse x < real (n::nat)"
390 using reals_Archimedean2 ..
391 hence "y * inverse x * x < real n * x"
392 using x_greater_zero by (simp add: mult_strict_right_mono)
393 hence "x * inverse x * y < x * real n"
394 by (simp add: ring_simps)
395 hence "y < real (n::nat) * x"
396 using x_not_zero by (simp add: ring_simps)
397 thus "\<exists>(n::nat). y < real n * x" ..
400 lemma reals_Archimedean6:
401 "0 \<le> r ==> \<exists>(n::nat). real (n - 1) \<le> r & r < real (n)"
402 apply (insert reals_Archimedean2 [of r], safe)
403 apply (subgoal_tac "\<exists>x::nat. r < real x \<and> (\<forall>y. r < real y \<longrightarrow> x \<le> y)", auto)
404 apply (rule_tac x = x in exI)
405 apply (case_tac x, simp)
406 apply (rename_tac x')
407 apply (drule_tac x = x' in spec, simp)
408 apply (rule_tac x="LEAST n. r < real n" in exI, safe)
409 apply (erule LeastI, erule Least_le)
412 lemma reals_Archimedean6a: "0 \<le> r ==> \<exists>n. real (n) \<le> r & r < real (Suc n)"
413 by (drule reals_Archimedean6) auto
415 lemma reals_Archimedean_6b_int:
416 "0 \<le> r ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
417 apply (drule reals_Archimedean6a, auto)
418 apply (rule_tac x = "int n" in exI)
419 apply (simp add: real_of_int_real_of_nat real_of_nat_Suc)
422 lemma reals_Archimedean_6c_int:
423 "r < 0 ==> \<exists>n::int. real n \<le> r & r < real (n+1)"
424 apply (rule reals_Archimedean_6b_int [of "-r", THEN exE], simp, auto)
426 apply (drule order_le_imp_less_or_eq, auto)
427 apply (rule_tac x = "- n - 1" in exI)
428 apply (rule_tac [2] x = "- n" in exI, auto)
432 subsection{*Density of the Rational Reals in the Reals*}
434 text{* This density proof is due to Stefan Richter and was ported by TN. The
435 original source is \emph{Real Analysis} by H.L. Royden.
436 It employs the Archimedean property of the reals. *}
438 lemma Rats_dense_in_nn_real: fixes x::real
439 assumes "0\<le>x" and "x<y" shows "\<exists>r \<in> \<rat>. x<r \<and> r<y"
441 from `x<y` have "0 < y-x" by simp
442 with reals_Archimedean obtain q::nat
443 where q: "inverse (real q) < y-x" and "0 < real q" by auto
444 def p \<equiv> "LEAST n. y \<le> real (Suc n)/real q"
445 from reals_Archimedean2 obtain n::nat where "y * real q < real n" by auto
446 with `0 < real q` have ex: "y \<le> real n/real q" (is "?P n")
447 by (simp add: pos_less_divide_eq[THEN sym])
448 also from assms have "\<not> y \<le> real (0::nat) / real q" by simp
449 ultimately have main: "(LEAST n. y \<le> real n/real q) = Suc p"
450 by (unfold p_def) (rule Least_Suc)
451 also from ex have "?P (LEAST x. ?P x)" by (rule LeastI)
452 ultimately have suc: "y \<le> real (Suc p) / real q" by simp
453 def r \<equiv> "real p/real q"
454 have "x = y-(y-x)" by simp
455 also from suc q have "\<dots> < real (Suc p)/real q - inverse (real q)" by arith
456 also have "\<dots> = real p / real q"
457 by (simp only: inverse_eq_divide real_diff_def real_of_nat_Suc
458 minus_divide_left add_divide_distrib[THEN sym]) simp
459 finally have "x<r" by (unfold r_def)
460 have "p<Suc p" .. also note main[THEN sym]
461 finally have "\<not> ?P p" by (rule not_less_Least)
462 hence "r<y" by (simp add: r_def)
463 from r_def have "r \<in> \<rat>" by simp
464 with `x<r` `r<y` show ?thesis by fast
467 theorem Rats_dense_in_real: fixes x y :: real
468 assumes "x<y" shows "\<exists>r \<in> \<rat>. x<r \<and> r<y"
470 from reals_Archimedean2 obtain n::nat where "-x < real n" by auto
471 hence "0 \<le> x + real n" by arith
472 also from `x<y` have "x + real n < y + real n" by arith
473 ultimately have "\<exists>r \<in> \<rat>. x + real n < r \<and> r < y + real n"
474 by(rule Rats_dense_in_nn_real)
475 then obtain r where "r \<in> \<rat>" and r2: "x + real n < r"
476 and r3: "r < y + real n"
478 have "r - real n = r + real (int n)/real (-1::int)" by simp
479 also from `r\<in>\<rat>` have "r + real (int n)/real (-1::int) \<in> \<rat>" by simp
480 also from r2 have "x < r - real n" by arith
481 moreover from r3 have "r - real n < y" by arith
482 ultimately show ?thesis by fast
486 subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
489 floor :: "real => int" where
490 [code del]: "floor r = (LEAST n::int. r < real (n+1))"
493 ceiling :: "real => int" where
494 "ceiling r = - floor (- r)"
497 floor ("\<lfloor>_\<rfloor>") and
498 ceiling ("\<lceil>_\<rceil>")
500 notation (HTML output)
501 floor ("\<lfloor>_\<rfloor>") and
502 ceiling ("\<lceil>_\<rceil>")
505 lemma number_of_less_real_of_int_iff [simp]:
506 "((number_of n) < real (m::int)) = (number_of n < m)"
508 apply (rule real_of_int_less_iff [THEN iffD1])
509 apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
512 lemma number_of_less_real_of_int_iff2 [simp]:
513 "(real (m::int) < (number_of n)) = (m < number_of n)"
515 apply (rule real_of_int_less_iff [THEN iffD1])
516 apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto)
519 lemma number_of_le_real_of_int_iff [simp]:
520 "((number_of n) \<le> real (m::int)) = (number_of n \<le> m)"
521 by (simp add: linorder_not_less [symmetric])
523 lemma number_of_le_real_of_int_iff2 [simp]:
524 "(real (m::int) \<le> (number_of n)) = (m \<le> number_of n)"
525 by (simp add: linorder_not_less [symmetric])
527 lemma floor_zero [simp]: "floor 0 = 0"
528 apply (simp add: floor_def del: real_of_int_add)
529 apply (rule Least_equality)
533 lemma floor_real_of_nat_zero [simp]: "floor (real (0::nat)) = 0"
536 lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n"
537 apply (simp only: floor_def)
538 apply (rule Least_equality)
539 apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst])
540 apply (drule_tac [2] real_of_int_less_iff [THEN iffD1])
544 lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
545 apply (simp only: floor_def)
546 apply (rule Least_equality)
547 apply (drule_tac [2] real_of_int_of_nat_eq [THEN ssubst])
548 apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst])
549 apply (drule_tac [2] real_of_int_less_iff [THEN iffD1])
553 lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
554 apply (simp only: floor_def)
555 apply (rule Least_equality)
559 lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
560 apply (simp only: floor_def)
561 apply (rule Least_equality)
562 apply (drule_tac [2] real_of_int_minus [THEN sym, THEN subst])
566 lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
567 apply (case_tac "r < 0")
568 apply (blast intro: reals_Archimedean_6c_int)
569 apply (simp only: linorder_not_less)
570 apply (blast intro: reals_Archimedean_6b_int reals_Archimedean_6c_int)
574 assumes a1: "real m \<le> r" and a2: "r < real n + 1"
575 shows "m \<le> (n::int)"
577 have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans)
578 also have "... = real (n + 1)" by simp
579 finally have "m < n + 1" by (simp only: real_of_int_less_iff)
580 thus ?thesis by arith
583 lemma real_of_int_floor_le [simp]: "real (floor r) \<le> r"
584 apply (simp add: floor_def Least_def)
585 apply (insert real_lb_ub_int [of r], safe)
590 lemma floor_mono: "x < y ==> floor x \<le> floor y"
591 apply (simp add: floor_def Least_def)
592 apply (insert real_lb_ub_int [of x])
593 apply (insert real_lb_ub_int [of y], safe)
595 apply (rule_tac [3] theI2)
598 apply (auto simp add: order_eq_iff int_le_real_less)
601 lemma floor_mono2: "x \<le> y ==> floor x \<le> floor y"
602 by (auto dest: order_le_imp_less_or_eq simp add: floor_mono)
604 lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \<le> x"
605 by (auto intro: lemma_floor)
607 lemma real_of_int_floor_cancel [simp]:
608 "(real (floor x) = x) = (\<exists>n::int. x = real n)"
609 apply (simp add: floor_def Least_def)
610 apply (insert real_lb_ub_int [of x], erule exE)
612 apply (auto intro: lemma_floor)
615 lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n"
616 apply (simp add: floor_def)
617 apply (rule Least_equality)
618 apply (auto intro: lemma_floor)
621 lemma floor_eq2: "[| real n \<le> x; x < real n + 1 |] ==> floor x = n"
622 apply (simp add: floor_def)
623 apply (rule Least_equality)
624 apply (auto intro: lemma_floor)
627 lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n"
628 apply (rule inj_int [THEN injD])
629 apply (simp add: real_of_nat_Suc)
630 apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"])
633 lemma floor_eq4: "[| real n \<le> x; x < real (Suc n) |] ==> nat(floor x) = n"
634 apply (drule order_le_imp_less_or_eq)
635 apply (auto intro: floor_eq3)
638 lemma floor_number_of_eq [simp]:
639 "floor(number_of n :: real) = (number_of n :: int)"
640 apply (subst real_number_of [symmetric])
641 apply (rule floor_real_of_int)
644 lemma floor_one [simp]: "floor 1 = 1"
647 apply (rule floor_real_of_int)
651 lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \<le> real(floor r)"
652 apply (simp add: floor_def Least_def)
653 apply (insert real_lb_ub_int [of r], safe)
655 apply (auto intro: lemma_floor)
658 lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)"
659 apply (simp add: floor_def Least_def)
660 apply (insert real_lb_ub_int [of r], safe)
662 apply (auto intro: lemma_floor)
665 lemma real_of_int_floor_add_one_ge [simp]: "r \<le> real(floor r) + 1"
666 apply (insert real_of_int_floor_ge_diff_one [of r])
667 apply (auto simp del: real_of_int_floor_ge_diff_one)
670 lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1"
671 apply (insert real_of_int_floor_gt_diff_one [of r])
672 apply (auto simp del: real_of_int_floor_gt_diff_one)
675 lemma le_floor: "real a <= x ==> a <= floor x"
676 apply (subgoal_tac "a < floor x + 1")
678 apply (subst real_of_int_less_iff [THEN sym])
680 apply (insert real_of_int_floor_add_one_gt [of x])
684 lemma real_le_floor: "a <= floor x ==> real a <= x"
685 apply (rule order_trans)
687 apply (rule real_of_int_floor_le)
688 apply (subst real_of_int_le_iff)
692 lemma le_floor_eq: "(a <= floor x) = (real a <= x)"
694 apply (erule real_le_floor)
695 apply (erule le_floor)
698 lemma le_floor_eq_number_of [simp]:
699 "(number_of n <= floor x) = (number_of n <= x)"
700 by (simp add: le_floor_eq)
702 lemma le_floor_eq_zero [simp]: "(0 <= floor x) = (0 <= x)"
703 by (simp add: le_floor_eq)
705 lemma le_floor_eq_one [simp]: "(1 <= floor x) = (1 <= x)"
706 by (simp add: le_floor_eq)
708 lemma floor_less_eq: "(floor x < a) = (x < real a)"
709 apply (subst linorder_not_le [THEN sym])+
711 apply (rule le_floor_eq)
714 lemma floor_less_eq_number_of [simp]:
715 "(floor x < number_of n) = (x < number_of n)"
716 by (simp add: floor_less_eq)
718 lemma floor_less_eq_zero [simp]: "(floor x < 0) = (x < 0)"
719 by (simp add: floor_less_eq)
721 lemma floor_less_eq_one [simp]: "(floor x < 1) = (x < 1)"
722 by (simp add: floor_less_eq)
724 lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)"
725 apply (insert le_floor_eq [of "a + 1" x])
729 lemma less_floor_eq_number_of [simp]:
730 "(number_of n < floor x) = (number_of n + 1 <= x)"
731 by (simp add: less_floor_eq)
733 lemma less_floor_eq_zero [simp]: "(0 < floor x) = (1 <= x)"
734 by (simp add: less_floor_eq)
736 lemma less_floor_eq_one [simp]: "(1 < floor x) = (2 <= x)"
737 by (simp add: less_floor_eq)
739 lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)"
740 apply (insert floor_less_eq [of x "a + 1"])
744 lemma floor_le_eq_number_of [simp]:
745 "(floor x <= number_of n) = (x < number_of n + 1)"
746 by (simp add: floor_le_eq)
748 lemma floor_le_eq_zero [simp]: "(floor x <= 0) = (x < 1)"
749 by (simp add: floor_le_eq)
751 lemma floor_le_eq_one [simp]: "(floor x <= 1) = (x < 2)"
752 by (simp add: floor_le_eq)
754 lemma floor_add [simp]: "floor (x + real a) = floor x + a"
755 apply (subst order_eq_iff)
758 apply (subgoal_tac "floor x + a < floor (x + real a) + 1")
760 apply (subst real_of_int_less_iff [THEN sym])
762 apply (subgoal_tac "x + real a < real(floor(x + real a)) + 1")
763 apply (subgoal_tac "real (floor x) <= x")
765 apply (rule real_of_int_floor_le)
766 apply (rule real_of_int_floor_add_one_gt)
767 apply (subgoal_tac "floor (x + real a) < floor x + a + 1")
769 apply (subst real_of_int_less_iff [THEN sym])
771 apply (subgoal_tac "real(floor(x + real a)) <= x + real a")
772 apply (subgoal_tac "x < real(floor x) + 1")
774 apply (rule real_of_int_floor_add_one_gt)
775 apply (rule real_of_int_floor_le)
778 lemma floor_add_number_of [simp]:
779 "floor (x + number_of n) = floor x + number_of n"
780 apply (subst floor_add [THEN sym])
784 lemma floor_add_one [simp]: "floor (x + 1) = floor x + 1"
785 apply (subst floor_add [THEN sym])
789 lemma floor_subtract [simp]: "floor (x - real a) = floor x - a"
790 apply (subst diff_minus)+
791 apply (subst real_of_int_minus [THEN sym])
792 apply (rule floor_add)
795 lemma floor_subtract_number_of [simp]: "floor (x - number_of n) =
796 floor x - number_of n"
797 apply (subst floor_subtract [THEN sym])
801 lemma floor_subtract_one [simp]: "floor (x - 1) = floor x - 1"
802 apply (subst floor_subtract [THEN sym])
806 lemma ceiling_zero [simp]: "ceiling 0 = 0"
807 by (simp add: ceiling_def)
809 lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n"
810 by (simp add: ceiling_def)
812 lemma ceiling_real_of_nat_zero [simp]: "ceiling (real (0::nat)) = 0"
815 lemma ceiling_floor [simp]: "ceiling (real (floor r)) = floor r"
816 by (simp add: ceiling_def)
818 lemma floor_ceiling [simp]: "floor (real (ceiling r)) = ceiling r"
819 by (simp add: ceiling_def)
821 lemma real_of_int_ceiling_ge [simp]: "r \<le> real (ceiling r)"
822 apply (simp add: ceiling_def)
823 apply (subst le_minus_iff, simp)
826 lemma ceiling_mono: "x < y ==> ceiling x \<le> ceiling y"
827 by (simp add: floor_mono ceiling_def)
829 lemma ceiling_mono2: "x \<le> y ==> ceiling x \<le> ceiling y"
830 by (simp add: floor_mono2 ceiling_def)
832 lemma real_of_int_ceiling_cancel [simp]:
833 "(real (ceiling x) = x) = (\<exists>n::int. x = real n)"
834 apply (auto simp add: ceiling_def)
835 apply (drule arg_cong [where f = uminus], auto)
836 apply (rule_tac x = "-n" in exI, auto)
839 lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1"
840 apply (simp add: ceiling_def)
841 apply (rule minus_equation_iff [THEN iffD1])
842 apply (simp add: floor_eq [where n = "-(n+1)"])
845 lemma ceiling_eq2: "[| real n < x; x \<le> real n + 1 |] ==> ceiling x = n + 1"
846 by (simp add: ceiling_def floor_eq2 [where n = "-(n+1)"])
848 lemma ceiling_eq3: "[| real n - 1 < x; x \<le> real n |] ==> ceiling x = n"
849 by (simp add: ceiling_def floor_eq2 [where n = "-n"])
851 lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n"
852 by (simp add: ceiling_def)
854 lemma ceiling_number_of_eq [simp]:
855 "ceiling (number_of n :: real) = (number_of n)"
856 apply (subst real_number_of [symmetric])
857 apply (rule ceiling_real_of_int)
860 lemma ceiling_one [simp]: "ceiling 1 = 1"
861 by (unfold ceiling_def, simp)
863 lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \<le> r"
864 apply (rule neg_le_iff_le [THEN iffD1])
865 apply (simp add: ceiling_def diff_minus)
868 lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \<le> r + 1"
869 apply (insert real_of_int_ceiling_diff_one_le [of r])
870 apply (simp del: real_of_int_ceiling_diff_one_le)
873 lemma ceiling_le: "x <= real a ==> ceiling x <= a"
874 apply (unfold ceiling_def)
875 apply (subgoal_tac "-a <= floor(- x)")
877 apply (rule le_floor)
881 lemma ceiling_le_real: "ceiling x <= a ==> x <= real a"
882 apply (unfold ceiling_def)
883 apply (subgoal_tac "real(- a) <= - x")
885 apply (rule real_le_floor)
889 lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)"
891 apply (erule ceiling_le_real)
892 apply (erule ceiling_le)
895 lemma ceiling_le_eq_number_of [simp]:
896 "(ceiling x <= number_of n) = (x <= number_of n)"
897 by (simp add: ceiling_le_eq)
899 lemma ceiling_le_zero_eq [simp]: "(ceiling x <= 0) = (x <= 0)"
900 by (simp add: ceiling_le_eq)
902 lemma ceiling_le_eq_one [simp]: "(ceiling x <= 1) = (x <= 1)"
903 by (simp add: ceiling_le_eq)
905 lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)"
906 apply (subst linorder_not_le [THEN sym])+
908 apply (rule ceiling_le_eq)
911 lemma less_ceiling_eq_number_of [simp]:
912 "(number_of n < ceiling x) = (number_of n < x)"
913 by (simp add: less_ceiling_eq)
915 lemma less_ceiling_eq_zero [simp]: "(0 < ceiling x) = (0 < x)"
916 by (simp add: less_ceiling_eq)
918 lemma less_ceiling_eq_one [simp]: "(1 < ceiling x) = (1 < x)"
919 by (simp add: less_ceiling_eq)
921 lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)"
922 apply (insert ceiling_le_eq [of x "a - 1"])
926 lemma ceiling_less_eq_number_of [simp]:
927 "(ceiling x < number_of n) = (x <= number_of n - 1)"
928 by (simp add: ceiling_less_eq)
930 lemma ceiling_less_eq_zero [simp]: "(ceiling x < 0) = (x <= -1)"
931 by (simp add: ceiling_less_eq)
933 lemma ceiling_less_eq_one [simp]: "(ceiling x < 1) = (x <= 0)"
934 by (simp add: ceiling_less_eq)
936 lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)"
937 apply (insert less_ceiling_eq [of "a - 1" x])
941 lemma le_ceiling_eq_number_of [simp]:
942 "(number_of n <= ceiling x) = (number_of n - 1 < x)"
943 by (simp add: le_ceiling_eq)
945 lemma le_ceiling_eq_zero [simp]: "(0 <= ceiling x) = (-1 < x)"
946 by (simp add: le_ceiling_eq)
948 lemma le_ceiling_eq_one [simp]: "(1 <= ceiling x) = (0 < x)"
949 by (simp add: le_ceiling_eq)
951 lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a"
952 apply (unfold ceiling_def, simp)
953 apply (subst real_of_int_minus [THEN sym])
954 apply (subst floor_add)
958 lemma ceiling_add_number_of [simp]: "ceiling (x + number_of n) =
959 ceiling x + number_of n"
960 apply (subst ceiling_add [THEN sym])
964 lemma ceiling_add_one [simp]: "ceiling (x + 1) = ceiling x + 1"
965 apply (subst ceiling_add [THEN sym])
969 lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a"
970 apply (subst diff_minus)+
971 apply (subst real_of_int_minus [THEN sym])
972 apply (rule ceiling_add)
975 lemma ceiling_subtract_number_of [simp]: "ceiling (x - number_of n) =
976 ceiling x - number_of n"
977 apply (subst ceiling_subtract [THEN sym])
981 lemma ceiling_subtract_one [simp]: "ceiling (x - 1) = ceiling x - 1"
982 apply (subst ceiling_subtract [THEN sym])
986 subsection {* Versions for the natural numbers *}
989 natfloor :: "real => nat" where
990 "natfloor x = nat(floor x)"
993 natceiling :: "real => nat" where
994 "natceiling x = nat(ceiling x)"
996 lemma natfloor_zero [simp]: "natfloor 0 = 0"
997 by (unfold natfloor_def, simp)
999 lemma natfloor_one [simp]: "natfloor 1 = 1"
1000 by (unfold natfloor_def, simp)
1002 lemma zero_le_natfloor [simp]: "0 <= natfloor x"
1003 by (unfold natfloor_def, simp)
1005 lemma natfloor_number_of_eq [simp]: "natfloor (number_of n) = number_of n"
1006 by (unfold natfloor_def, simp)
1008 lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n"
1009 by (unfold natfloor_def, simp)
1011 lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x"
1012 by (unfold natfloor_def, simp)
1014 lemma natfloor_neg: "x <= 0 ==> natfloor x = 0"
1015 apply (unfold natfloor_def)
1016 apply (subgoal_tac "floor x <= floor 0")
1018 apply (erule floor_mono2)
1021 lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y"
1022 apply (case_tac "0 <= x")
1023 apply (subst natfloor_def)+
1024 apply (subst nat_le_eq_zle)
1026 apply (erule floor_mono2)
1027 apply (subst natfloor_neg)
1032 lemma le_natfloor: "real x <= a ==> x <= natfloor a"
1033 apply (unfold natfloor_def)
1034 apply (subst nat_int [THEN sym])
1035 apply (subst nat_le_eq_zle)
1037 apply (rule le_floor)
1041 lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)"
1043 apply (rule order_trans)
1045 apply (erule real_natfloor_le)
1046 apply (subst real_of_nat_le_iff)
1048 apply (erule le_natfloor)
1051 lemma le_natfloor_eq_number_of [simp]:
1052 "~ neg((number_of n)::int) ==> 0 <= x ==>
1053 (number_of n <= natfloor x) = (number_of n <= x)"
1054 apply (subst le_natfloor_eq, assumption)
1058 lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)"
1059 apply (case_tac "0 <= x")
1060 apply (subst le_natfloor_eq, assumption, simp)
1062 apply (subgoal_tac "natfloor x <= natfloor 0")
1064 apply (rule natfloor_mono)
1069 lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n"
1070 apply (unfold natfloor_def)
1071 apply (subst nat_int [THEN sym]);back;
1072 apply (subst eq_nat_nat_iff)
1075 apply (rule floor_eq2)
1079 lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1"
1080 apply (case_tac "0 <= x")
1081 apply (unfold natfloor_def)
1086 lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)"
1087 apply (simp add: compare_rls)
1088 apply (rule real_natfloor_add_one_gt)
1091 lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n"
1092 apply (subgoal_tac "z < real(natfloor z) + 1")
1094 apply (rule real_natfloor_add_one_gt)
1097 lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a"
1098 apply (unfold natfloor_def)
1099 apply (subgoal_tac "real a = real (int a)")
1100 apply (erule ssubst)
1101 apply (simp add: nat_add_distrib del: real_of_int_of_nat_eq)
1105 lemma natfloor_add_number_of [simp]:
1106 "~neg ((number_of n)::int) ==> 0 <= x ==>
1107 natfloor (x + number_of n) = natfloor x + number_of n"
1108 apply (subst natfloor_add [THEN sym])
1112 lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1"
1113 apply (subst natfloor_add [THEN sym])
1118 lemma natfloor_subtract [simp]: "real a <= x ==>
1119 natfloor(x - real a) = natfloor x - a"
1120 apply (unfold natfloor_def)
1121 apply (subgoal_tac "real a = real (int a)")
1122 apply (erule ssubst)
1123 apply (simp del: real_of_int_of_nat_eq)
1127 lemma natceiling_zero [simp]: "natceiling 0 = 0"
1128 by (unfold natceiling_def, simp)
1130 lemma natceiling_one [simp]: "natceiling 1 = 1"
1131 by (unfold natceiling_def, simp)
1133 lemma zero_le_natceiling [simp]: "0 <= natceiling x"
1134 by (unfold natceiling_def, simp)
1136 lemma natceiling_number_of_eq [simp]: "natceiling (number_of n) = number_of n"
1137 by (unfold natceiling_def, simp)
1139 lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n"
1140 by (unfold natceiling_def, simp)
1142 lemma real_natceiling_ge: "x <= real(natceiling x)"
1143 apply (unfold natceiling_def)
1144 apply (case_tac "x < 0")
1146 apply (subst real_nat_eq_real)
1147 apply (subgoal_tac "ceiling 0 <= ceiling x")
1149 apply (rule ceiling_mono2)
1154 lemma natceiling_neg: "x <= 0 ==> natceiling x = 0"
1155 apply (unfold natceiling_def)
1159 lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y"
1160 apply (case_tac "0 <= x")
1161 apply (subst natceiling_def)+
1162 apply (subst nat_le_eq_zle)
1164 apply (subgoal_tac "real (0::int) <= real(ceiling y)")
1166 apply (rule order_trans)
1168 apply (erule order_trans)
1170 apply (erule ceiling_mono2)
1171 apply (subst natceiling_neg)
1175 lemma natceiling_le: "x <= real a ==> natceiling x <= a"
1176 apply (unfold natceiling_def)
1177 apply (case_tac "x < 0")
1179 apply (subst nat_int [THEN sym]);back;
1180 apply (subst nat_le_eq_zle)
1182 apply (rule ceiling_le)
1186 lemma natceiling_le_eq: "0 <= x ==> (natceiling x <= a) = (x <= real a)"
1188 apply (rule order_trans)
1189 apply (rule real_natceiling_ge)
1190 apply (subst real_of_nat_le_iff)
1192 apply (erule natceiling_le)
1195 lemma natceiling_le_eq_number_of [simp]:
1196 "~ neg((number_of n)::int) ==> 0 <= x ==>
1197 (natceiling x <= number_of n) = (x <= number_of n)"
1198 apply (subst natceiling_le_eq, assumption)
1202 lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)"
1203 apply (case_tac "0 <= x")
1204 apply (subst natceiling_le_eq)
1207 apply (subst natceiling_neg)
1212 lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1"
1213 apply (unfold natceiling_def)
1214 apply (simplesubst nat_int [THEN sym]) back back
1215 apply (subgoal_tac "nat(int n) + 1 = nat(int n + 1)")
1216 apply (erule ssubst)
1217 apply (subst eq_nat_nat_iff)
1218 apply (subgoal_tac "ceiling 0 <= ceiling x")
1220 apply (rule ceiling_mono2)
1223 apply (rule ceiling_eq2)
1225 apply (subst nat_add_distrib)
1229 lemma natceiling_add [simp]: "0 <= x ==>
1230 natceiling (x + real a) = natceiling x + a"
1231 apply (unfold natceiling_def)
1232 apply (subgoal_tac "real a = real (int a)")
1233 apply (erule ssubst)
1234 apply (simp del: real_of_int_of_nat_eq)
1235 apply (subst nat_add_distrib)
1236 apply (subgoal_tac "0 = ceiling 0")
1237 apply (erule ssubst)
1238 apply (erule ceiling_mono2)
1242 lemma natceiling_add_number_of [simp]:
1243 "~ neg ((number_of n)::int) ==> 0 <= x ==>
1244 natceiling (x + number_of n) = natceiling x + number_of n"
1245 apply (subst natceiling_add [THEN sym])
1249 lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1"
1250 apply (subst natceiling_add [THEN sym])
1255 lemma natceiling_subtract [simp]: "real a <= x ==>
1256 natceiling(x - real a) = natceiling x - a"
1257 apply (unfold natceiling_def)
1258 apply (subgoal_tac "real a = real (int a)")
1259 apply (erule ssubst)
1260 apply (simp del: real_of_int_of_nat_eq)
1264 lemma natfloor_div_nat: "1 <= x ==> y > 0 ==>
1265 natfloor (x / real y) = natfloor x div y"
1267 assume "1 <= (x::real)" and "(y::nat) > 0"
1268 have "natfloor x = (natfloor x) div y * y + (natfloor x) mod y"
1270 then have a: "real(natfloor x) = real ((natfloor x) div y) * real y +
1271 real((natfloor x) mod y)"
1272 by (simp only: real_of_nat_add [THEN sym] real_of_nat_mult [THEN sym])
1273 have "x = real(natfloor x) + (x - real(natfloor x))"
1275 then have "x = real ((natfloor x) div y) * real y +
1276 real((natfloor x) mod y) + (x - real(natfloor x))"
1278 then have "x / real y = ... / real y"
1280 also have "... = real((natfloor x) div y) + real((natfloor x) mod y) /
1281 real y + (x - real(natfloor x)) / real y"
1282 by (auto simp add: ring_simps add_divide_distrib
1283 diff_divide_distrib prems)
1284 finally have "natfloor (x / real y) = natfloor(...)" by simp
1285 also have "... = natfloor(real((natfloor x) mod y) /
1286 real y + (x - real(natfloor x)) / real y + real((natfloor x) div y))"
1287 by (simp add: add_ac)
1288 also have "... = natfloor(real((natfloor x) mod y) /
1289 real y + (x - real(natfloor x)) / real y) + (natfloor x) div y"
1290 apply (rule natfloor_add)
1291 apply (rule add_nonneg_nonneg)
1292 apply (rule divide_nonneg_pos)
1294 apply (simp add: prems)
1295 apply (rule divide_nonneg_pos)
1296 apply (simp add: compare_rls)
1297 apply (rule real_natfloor_le)
1298 apply (insert prems, auto)
1300 also have "natfloor(real((natfloor x) mod y) /
1301 real y + (x - real(natfloor x)) / real y) = 0"
1302 apply (rule natfloor_eq)
1304 apply (rule add_nonneg_nonneg)
1305 apply (rule divide_nonneg_pos)
1307 apply (force simp add: prems)
1308 apply (rule divide_nonneg_pos)
1309 apply (simp add: compare_rls)
1310 apply (rule real_natfloor_le)
1311 apply (auto simp add: prems)
1312 apply (insert prems, arith)
1313 apply (simp add: add_divide_distrib [THEN sym])
1314 apply (subgoal_tac "real y = real y - 1 + 1")
1315 apply (erule ssubst)
1316 apply (rule add_le_less_mono)
1317 apply (simp add: compare_rls)
1318 apply (subgoal_tac "real(natfloor x mod y) + 1 =
1319 real(natfloor x mod y + 1)")
1320 apply (erule ssubst)
1321 apply (subst real_of_nat_le_iff)
1322 apply (subgoal_tac "natfloor x mod y < y")
1324 apply (rule mod_less_divisor)
1326 apply (simp add: compare_rls)
1327 apply (subst add_commute)
1328 apply (rule real_natfloor_add_one_gt)
1330 finally show ?thesis by simp