hoelzl@51523: (* Title: HOL/Real.thy hoelzl@51523: Author: Jacques D. Fleuriot, University of Edinburgh, 1998 hoelzl@51523: Author: Larry Paulson, University of Cambridge hoelzl@51523: Author: Jeremy Avigad, Carnegie Mellon University hoelzl@51523: Author: Florian Zuleger, Johannes Hoelzl, and Simon Funke, TU Muenchen hoelzl@51523: Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 hoelzl@51523: Construction of Cauchy Reals by Brian Huffman, 2010 hoelzl@51523: *) hoelzl@51523: hoelzl@51523: header {* Development of the Reals using Cauchy Sequences *} hoelzl@51523: hoelzl@51523: theory Real hoelzl@51773: imports Rat Conditionally_Complete_Lattices hoelzl@51523: begin hoelzl@51523: hoelzl@51523: text {* hoelzl@51523: This theory contains a formalization of the real numbers as hoelzl@51523: equivalence classes of Cauchy sequences of rationals. See hoelzl@51523: @{file "~~/src/HOL/ex/Dedekind_Real.thy"} for an alternative hoelzl@51523: construction using Dedekind cuts. hoelzl@51523: *} hoelzl@51523: hoelzl@51523: subsection {* Preliminary lemmas *} hoelzl@51523: hoelzl@51523: lemma add_diff_add: hoelzl@51523: fixes a b c d :: "'a::ab_group_add" hoelzl@51523: shows "(a + c) - (b + d) = (a - b) + (c - d)" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma minus_diff_minus: hoelzl@51523: fixes a b :: "'a::ab_group_add" hoelzl@51523: shows "- a - - b = - (a - b)" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma mult_diff_mult: hoelzl@51523: fixes x y a b :: "'a::ring" hoelzl@51523: shows "(x * y - a * b) = x * (y - b) + (x - a) * b" hoelzl@51523: by (simp add: algebra_simps) hoelzl@51523: hoelzl@51523: lemma inverse_diff_inverse: hoelzl@51523: fixes a b :: "'a::division_ring" hoelzl@51523: assumes "a \ 0" and "b \ 0" hoelzl@51523: shows "inverse a - inverse b = - (inverse a * (a - b) * inverse b)" hoelzl@51523: using assms by (simp add: algebra_simps) hoelzl@51523: hoelzl@51523: lemma obtain_pos_sum: hoelzl@51523: fixes r :: rat assumes r: "0 < r" hoelzl@51523: obtains s t where "0 < s" and "0 < t" and "r = s + t" hoelzl@51523: proof hoelzl@51523: from r show "0 < r/2" by simp hoelzl@51523: from r show "0 < r/2" by simp hoelzl@51523: show "r = r/2 + r/2" by simp hoelzl@51523: qed hoelzl@51523: hoelzl@51523: subsection {* Sequences that converge to zero *} hoelzl@51523: hoelzl@51523: definition hoelzl@51523: vanishes :: "(nat \ rat) \ bool" hoelzl@51523: where hoelzl@51523: "vanishes X = (\r>0. \k. \n\k. \X n\ < r)" hoelzl@51523: hoelzl@51523: lemma vanishesI: "(\r. 0 < r \ \k. \n\k. \X n\ < r) \ vanishes X" hoelzl@51523: unfolding vanishes_def by simp hoelzl@51523: hoelzl@51523: lemma vanishesD: "\vanishes X; 0 < r\ \ \k. \n\k. \X n\ < r" hoelzl@51523: unfolding vanishes_def by simp hoelzl@51523: hoelzl@51523: lemma vanishes_const [simp]: "vanishes (\n. c) \ c = 0" hoelzl@51523: unfolding vanishes_def hoelzl@51523: apply (cases "c = 0", auto) hoelzl@51523: apply (rule exI [where x="\c\"], auto) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma vanishes_minus: "vanishes X \ vanishes (\n. - X n)" hoelzl@51523: unfolding vanishes_def by simp hoelzl@51523: hoelzl@51523: lemma vanishes_add: hoelzl@51523: assumes X: "vanishes X" and Y: "vanishes Y" hoelzl@51523: shows "vanishes (\n. X n + Y n)" hoelzl@51523: proof (rule vanishesI) hoelzl@51523: fix r :: rat assume "0 < r" hoelzl@51523: then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" hoelzl@51523: by (rule obtain_pos_sum) hoelzl@51523: obtain i where i: "\n\i. \X n\ < s" hoelzl@51523: using vanishesD [OF X s] .. hoelzl@51523: obtain j where j: "\n\j. \Y n\ < t" hoelzl@51523: using vanishesD [OF Y t] .. hoelzl@51523: have "\n\max i j. \X n + Y n\ < r" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix n assume n: "i \ n" "j \ n" hoelzl@51523: have "\X n + Y n\ \ \X n\ + \Y n\" by (rule abs_triangle_ineq) hoelzl@51523: also have "\ < s + t" by (simp add: add_strict_mono i j n) hoelzl@51523: finally show "\X n + Y n\ < r" unfolding r . hoelzl@51523: qed hoelzl@51523: thus "\k. \n\k. \X n + Y n\ < r" .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma vanishes_diff: hoelzl@51523: assumes X: "vanishes X" and Y: "vanishes Y" hoelzl@51523: shows "vanishes (\n. X n - Y n)" hoelzl@51523: unfolding diff_minus by (intro vanishes_add vanishes_minus X Y) hoelzl@51523: hoelzl@51523: lemma vanishes_mult_bounded: hoelzl@51523: assumes X: "\a>0. \n. \X n\ < a" hoelzl@51523: assumes Y: "vanishes (\n. Y n)" hoelzl@51523: shows "vanishes (\n. X n * Y n)" hoelzl@51523: proof (rule vanishesI) hoelzl@51523: fix r :: rat assume r: "0 < r" hoelzl@51523: obtain a where a: "0 < a" "\n. \X n\ < a" hoelzl@51523: using X by fast hoelzl@51523: obtain b where b: "0 < b" "r = a * b" hoelzl@51523: proof hoelzl@51523: show "0 < r / a" using r a by (simp add: divide_pos_pos) hoelzl@51523: show "r = a * (r / a)" using a by simp hoelzl@51523: qed hoelzl@51523: obtain k where k: "\n\k. \Y n\ < b" hoelzl@51523: using vanishesD [OF Y b(1)] .. hoelzl@51523: have "\n\k. \X n * Y n\ < r" hoelzl@51523: by (simp add: b(2) abs_mult mult_strict_mono' a k) hoelzl@51523: thus "\k. \n\k. \X n * Y n\ < r" .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: subsection {* Cauchy sequences *} hoelzl@51523: hoelzl@51523: definition hoelzl@51523: cauchy :: "(nat \ rat) \ bool" hoelzl@51523: where hoelzl@51523: "cauchy X \ (\r>0. \k. \m\k. \n\k. \X m - X n\ < r)" hoelzl@51523: hoelzl@51523: lemma cauchyI: hoelzl@51523: "(\r. 0 < r \ \k. \m\k. \n\k. \X m - X n\ < r) \ cauchy X" hoelzl@51523: unfolding cauchy_def by simp hoelzl@51523: hoelzl@51523: lemma cauchyD: hoelzl@51523: "\cauchy X; 0 < r\ \ \k. \m\k. \n\k. \X m - X n\ < r" hoelzl@51523: unfolding cauchy_def by simp hoelzl@51523: hoelzl@51523: lemma cauchy_const [simp]: "cauchy (\n. x)" hoelzl@51523: unfolding cauchy_def by simp hoelzl@51523: hoelzl@51523: lemma cauchy_add [simp]: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "cauchy (\n. X n + Y n)" hoelzl@51523: proof (rule cauchyI) hoelzl@51523: fix r :: rat assume "0 < r" hoelzl@51523: then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" hoelzl@51523: by (rule obtain_pos_sum) hoelzl@51523: obtain i where i: "\m\i. \n\i. \X m - X n\ < s" hoelzl@51523: using cauchyD [OF X s] .. hoelzl@51523: obtain j where j: "\m\j. \n\j. \Y m - Y n\ < t" hoelzl@51523: using cauchyD [OF Y t] .. hoelzl@51523: have "\m\max i j. \n\max i j. \(X m + Y m) - (X n + Y n)\ < r" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" hoelzl@51523: have "\(X m + Y m) - (X n + Y n)\ \ \X m - X n\ + \Y m - Y n\" hoelzl@51523: unfolding add_diff_add by (rule abs_triangle_ineq) hoelzl@51523: also have "\ < s + t" hoelzl@51523: by (rule add_strict_mono, simp_all add: i j *) hoelzl@51523: finally show "\(X m + Y m) - (X n + Y n)\ < r" unfolding r . hoelzl@51523: qed hoelzl@51523: thus "\k. \m\k. \n\k. \(X m + Y m) - (X n + Y n)\ < r" .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma cauchy_minus [simp]: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: shows "cauchy (\n. - X n)" hoelzl@51523: using assms unfolding cauchy_def hoelzl@51523: unfolding minus_diff_minus abs_minus_cancel . hoelzl@51523: hoelzl@51523: lemma cauchy_diff [simp]: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "cauchy (\n. X n - Y n)" hoelzl@51523: using assms unfolding diff_minus by simp hoelzl@51523: hoelzl@51523: lemma cauchy_imp_bounded: hoelzl@51523: assumes "cauchy X" shows "\b>0. \n. \X n\ < b" hoelzl@51523: proof - hoelzl@51523: obtain k where k: "\m\k. \n\k. \X m - X n\ < 1" hoelzl@51523: using cauchyD [OF assms zero_less_one] .. hoelzl@51523: show "\b>0. \n. \X n\ < b" hoelzl@51523: proof (intro exI conjI allI) hoelzl@51523: have "0 \ \X 0\" by simp hoelzl@51523: also have "\X 0\ \ Max (abs ` X ` {..k})" by simp hoelzl@51523: finally have "0 \ Max (abs ` X ` {..k})" . hoelzl@51523: thus "0 < Max (abs ` X ` {..k}) + 1" by simp hoelzl@51523: next hoelzl@51523: fix n :: nat hoelzl@51523: show "\X n\ < Max (abs ` X ` {..k}) + 1" hoelzl@51523: proof (rule linorder_le_cases) hoelzl@51523: assume "n \ k" hoelzl@51523: hence "\X n\ \ Max (abs ` X ` {..k})" by simp hoelzl@51523: thus "\X n\ < Max (abs ` X ` {..k}) + 1" by simp hoelzl@51523: next hoelzl@51523: assume "k \ n" hoelzl@51523: have "\X n\ = \X k + (X n - X k)\" by simp hoelzl@51523: also have "\X k + (X n - X k)\ \ \X k\ + \X n - X k\" hoelzl@51523: by (rule abs_triangle_ineq) hoelzl@51523: also have "\ < Max (abs ` X ` {..k}) + 1" hoelzl@51523: by (rule add_le_less_mono, simp, simp add: k `k \ n`) hoelzl@51523: finally show "\X n\ < Max (abs ` X ` {..k}) + 1" . hoelzl@51523: qed hoelzl@51523: qed hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma cauchy_mult [simp]: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "cauchy (\n. X n * Y n)" hoelzl@51523: proof (rule cauchyI) hoelzl@51523: fix r :: rat assume "0 < r" hoelzl@51523: then obtain u v where u: "0 < u" and v: "0 < v" and "r = u + v" hoelzl@51523: by (rule obtain_pos_sum) hoelzl@51523: obtain a where a: "0 < a" "\n. \X n\ < a" hoelzl@51523: using cauchy_imp_bounded [OF X] by fast hoelzl@51523: obtain b where b: "0 < b" "\n. \Y n\ < b" hoelzl@51523: using cauchy_imp_bounded [OF Y] by fast hoelzl@51523: obtain s t where s: "0 < s" and t: "0 < t" and r: "r = a * t + s * b" hoelzl@51523: proof hoelzl@51523: show "0 < v/b" using v b(1) by (rule divide_pos_pos) hoelzl@51523: show "0 < u/a" using u a(1) by (rule divide_pos_pos) hoelzl@51523: show "r = a * (u/a) + (v/b) * b" hoelzl@51523: using a(1) b(1) `r = u + v` by simp hoelzl@51523: qed hoelzl@51523: obtain i where i: "\m\i. \n\i. \X m - X n\ < s" hoelzl@51523: using cauchyD [OF X s] .. hoelzl@51523: obtain j where j: "\m\j. \n\j. \Y m - Y n\ < t" hoelzl@51523: using cauchyD [OF Y t] .. hoelzl@51523: have "\m\max i j. \n\max i j. \X m * Y m - X n * Y n\ < r" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" hoelzl@51523: have "\X m * Y m - X n * Y n\ = \X m * (Y m - Y n) + (X m - X n) * Y n\" hoelzl@51523: unfolding mult_diff_mult .. hoelzl@51523: also have "\ \ \X m * (Y m - Y n)\ + \(X m - X n) * Y n\" hoelzl@51523: by (rule abs_triangle_ineq) hoelzl@51523: also have "\ = \X m\ * \Y m - Y n\ + \X m - X n\ * \Y n\" hoelzl@51523: unfolding abs_mult .. hoelzl@51523: also have "\ < a * t + s * b" hoelzl@51523: by (simp_all add: add_strict_mono mult_strict_mono' a b i j *) hoelzl@51523: finally show "\X m * Y m - X n * Y n\ < r" unfolding r . hoelzl@51523: qed hoelzl@51523: thus "\k. \m\k. \n\k. \X m * Y m - X n * Y n\ < r" .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma cauchy_not_vanishes_cases: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: assumes nz: "\ vanishes X" hoelzl@51523: shows "\b>0. \k. (\n\k. b < - X n) \ (\n\k. b < X n)" hoelzl@51523: proof - hoelzl@51523: obtain r where "0 < r" and r: "\k. \n\k. r \ \X n\" hoelzl@51523: using nz unfolding vanishes_def by (auto simp add: not_less) hoelzl@51523: obtain s t where s: "0 < s" and t: "0 < t" and "r = s + t" hoelzl@51523: using `0 < r` by (rule obtain_pos_sum) hoelzl@51523: obtain i where i: "\m\i. \n\i. \X m - X n\ < s" hoelzl@51523: using cauchyD [OF X s] .. hoelzl@51523: obtain k where "i \ k" and "r \ \X k\" hoelzl@51523: using r by fast hoelzl@51523: have k: "\n\k. \X n - X k\ < s" hoelzl@51523: using i `i \ k` by auto hoelzl@51523: have "X k \ - r \ r \ X k" hoelzl@51523: using `r \ \X k\` by auto hoelzl@51523: hence "(\n\k. t < - X n) \ (\n\k. t < X n)" hoelzl@51523: unfolding `r = s + t` using k by auto hoelzl@51523: hence "\k. (\n\k. t < - X n) \ (\n\k. t < X n)" .. hoelzl@51523: thus "\t>0. \k. (\n\k. t < - X n) \ (\n\k. t < X n)" hoelzl@51523: using t by auto hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma cauchy_not_vanishes: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: assumes nz: "\ vanishes X" hoelzl@51523: shows "\b>0. \k. \n\k. b < \X n\" hoelzl@51523: using cauchy_not_vanishes_cases [OF assms] hoelzl@51523: by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto) hoelzl@51523: hoelzl@51523: lemma cauchy_inverse [simp]: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: assumes nz: "\ vanishes X" hoelzl@51523: shows "cauchy (\n. inverse (X n))" hoelzl@51523: proof (rule cauchyI) hoelzl@51523: fix r :: rat assume "0 < r" hoelzl@51523: obtain b i where b: "0 < b" and i: "\n\i. b < \X n\" hoelzl@51523: using cauchy_not_vanishes [OF X nz] by fast hoelzl@51523: from b i have nz: "\n\i. X n \ 0" by auto hoelzl@51523: obtain s where s: "0 < s" and r: "r = inverse b * s * inverse b" hoelzl@51523: proof hoelzl@51523: show "0 < b * r * b" hoelzl@51523: by (simp add: `0 < r` b mult_pos_pos) hoelzl@51523: show "r = inverse b * (b * r * b) * inverse b" hoelzl@51523: using b by simp hoelzl@51523: qed hoelzl@51523: obtain j where j: "\m\j. \n\j. \X m - X n\ < s" hoelzl@51523: using cauchyD [OF X s] .. hoelzl@51523: have "\m\max i j. \n\max i j. \inverse (X m) - inverse (X n)\ < r" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" hoelzl@51523: have "\inverse (X m) - inverse (X n)\ = hoelzl@51523: inverse \X m\ * \X m - X n\ * inverse \X n\" hoelzl@51523: by (simp add: inverse_diff_inverse nz * abs_mult) hoelzl@51523: also have "\ < inverse b * s * inverse b" hoelzl@51523: by (simp add: mult_strict_mono less_imp_inverse_less hoelzl@51523: mult_pos_pos i j b * s) hoelzl@51523: finally show "\inverse (X m) - inverse (X n)\ < r" unfolding r . hoelzl@51523: qed hoelzl@51523: thus "\k. \m\k. \n\k. \inverse (X m) - inverse (X n)\ < r" .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma vanishes_diff_inverse: hoelzl@51523: assumes X: "cauchy X" "\ vanishes X" hoelzl@51523: assumes Y: "cauchy Y" "\ vanishes Y" hoelzl@51523: assumes XY: "vanishes (\n. X n - Y n)" hoelzl@51523: shows "vanishes (\n. inverse (X n) - inverse (Y n))" hoelzl@51523: proof (rule vanishesI) hoelzl@51523: fix r :: rat assume r: "0 < r" hoelzl@51523: obtain a i where a: "0 < a" and i: "\n\i. a < \X n\" hoelzl@51523: using cauchy_not_vanishes [OF X] by fast hoelzl@51523: obtain b j where b: "0 < b" and j: "\n\j. b < \Y n\" hoelzl@51523: using cauchy_not_vanishes [OF Y] by fast hoelzl@51523: obtain s where s: "0 < s" and "inverse a * s * inverse b = r" hoelzl@51523: proof hoelzl@51523: show "0 < a * r * b" hoelzl@51523: using a r b by (simp add: mult_pos_pos) hoelzl@51523: show "inverse a * (a * r * b) * inverse b = r" hoelzl@51523: using a r b by simp hoelzl@51523: qed hoelzl@51523: obtain k where k: "\n\k. \X n - Y n\ < s" hoelzl@51523: using vanishesD [OF XY s] .. hoelzl@51523: have "\n\max (max i j) k. \inverse (X n) - inverse (Y n)\ < r" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix n assume n: "i \ n" "j \ n" "k \ n" hoelzl@51523: have "X n \ 0" and "Y n \ 0" hoelzl@51523: using i j a b n by auto hoelzl@51523: hence "\inverse (X n) - inverse (Y n)\ = hoelzl@51523: inverse \X n\ * \X n - Y n\ * inverse \Y n\" hoelzl@51523: by (simp add: inverse_diff_inverse abs_mult) hoelzl@51523: also have "\ < inverse a * s * inverse b" hoelzl@51523: apply (intro mult_strict_mono' less_imp_inverse_less) hoelzl@51523: apply (simp_all add: a b i j k n mult_nonneg_nonneg) hoelzl@51523: done hoelzl@51523: also note `inverse a * s * inverse b = r` hoelzl@51523: finally show "\inverse (X n) - inverse (Y n)\ < r" . hoelzl@51523: qed hoelzl@51523: thus "\k. \n\k. \inverse (X n) - inverse (Y n)\ < r" .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: subsection {* Equivalence relation on Cauchy sequences *} hoelzl@51523: hoelzl@51523: definition realrel :: "(nat \ rat) \ (nat \ rat) \ bool" hoelzl@51523: where "realrel = (\X Y. cauchy X \ cauchy Y \ vanishes (\n. X n - Y n))" hoelzl@51523: hoelzl@51523: lemma realrelI [intro?]: hoelzl@51523: assumes "cauchy X" and "cauchy Y" and "vanishes (\n. X n - Y n)" hoelzl@51523: shows "realrel X Y" hoelzl@51523: using assms unfolding realrel_def by simp hoelzl@51523: hoelzl@51523: lemma realrel_refl: "cauchy X \ realrel X X" hoelzl@51523: unfolding realrel_def by simp hoelzl@51523: hoelzl@51523: lemma symp_realrel: "symp realrel" hoelzl@51523: unfolding realrel_def hoelzl@51523: by (rule sympI, clarify, drule vanishes_minus, simp) hoelzl@51523: hoelzl@51523: lemma transp_realrel: "transp realrel" hoelzl@51523: unfolding realrel_def hoelzl@51523: apply (rule transpI, clarify) hoelzl@51523: apply (drule (1) vanishes_add) hoelzl@51523: apply (simp add: algebra_simps) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma part_equivp_realrel: "part_equivp realrel" hoelzl@51523: by (fast intro: part_equivpI symp_realrel transp_realrel hoelzl@51523: realrel_refl cauchy_const) hoelzl@51523: hoelzl@51523: subsection {* The field of real numbers *} hoelzl@51523: hoelzl@51523: quotient_type real = "nat \ rat" / partial: realrel hoelzl@51523: morphisms rep_real Real hoelzl@51523: by (rule part_equivp_realrel) hoelzl@51523: hoelzl@51523: lemma cr_real_eq: "pcr_real = (\x y. cauchy x \ Real x = y)" hoelzl@51523: unfolding real.pcr_cr_eq cr_real_def realrel_def by auto hoelzl@51523: hoelzl@51523: lemma Real_induct [induct type: real]: (* TODO: generate automatically *) hoelzl@51523: assumes "\X. cauchy X \ P (Real X)" shows "P x" hoelzl@51523: proof (induct x) hoelzl@51523: case (1 X) hoelzl@51523: hence "cauchy X" by (simp add: realrel_def) hoelzl@51523: thus "P (Real X)" by (rule assms) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma eq_Real: hoelzl@51523: "cauchy X \ cauchy Y \ Real X = Real Y \ vanishes (\n. X n - Y n)" hoelzl@51523: using real.rel_eq_transfer hoelzl@51523: unfolding real.pcr_cr_eq cr_real_def fun_rel_def realrel_def by simp hoelzl@51523: hoelzl@51523: declare real.forall_transfer [transfer_rule del] hoelzl@51523: hoelzl@51523: lemma forall_real_transfer [transfer_rule]: (* TODO: generate automatically *) hoelzl@51523: "(fun_rel (fun_rel pcr_real op =) op =) hoelzl@51523: (transfer_bforall cauchy) transfer_forall" hoelzl@51523: using real.forall_transfer hoelzl@51523: by (simp add: realrel_def) hoelzl@51523: hoelzl@51523: instantiation real :: field_inverse_zero hoelzl@51523: begin hoelzl@51523: hoelzl@51523: lift_definition zero_real :: "real" is "\n. 0" hoelzl@51523: by (simp add: realrel_refl) hoelzl@51523: hoelzl@51523: lift_definition one_real :: "real" is "\n. 1" hoelzl@51523: by (simp add: realrel_refl) hoelzl@51523: hoelzl@51523: lift_definition plus_real :: "real \ real \ real" is "\X Y n. X n + Y n" hoelzl@51523: unfolding realrel_def add_diff_add hoelzl@51523: by (simp only: cauchy_add vanishes_add simp_thms) hoelzl@51523: hoelzl@51523: lift_definition uminus_real :: "real \ real" is "\X n. - X n" hoelzl@51523: unfolding realrel_def minus_diff_minus hoelzl@51523: by (simp only: cauchy_minus vanishes_minus simp_thms) hoelzl@51523: hoelzl@51523: lift_definition times_real :: "real \ real \ real" is "\X Y n. X n * Y n" hoelzl@51523: unfolding realrel_def mult_diff_mult hoelzl@51523: by (subst (4) mult_commute, simp only: cauchy_mult vanishes_add hoelzl@51523: vanishes_mult_bounded cauchy_imp_bounded simp_thms) hoelzl@51523: hoelzl@51523: lift_definition inverse_real :: "real \ real" hoelzl@51523: is "\X. if vanishes X then (\n. 0) else (\n. inverse (X n))" hoelzl@51523: proof - hoelzl@51523: fix X Y assume "realrel X Y" hoelzl@51523: hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\n. X n - Y n)" hoelzl@51523: unfolding realrel_def by simp_all hoelzl@51523: have "vanishes X \ vanishes Y" hoelzl@51523: proof hoelzl@51523: assume "vanishes X" hoelzl@51523: from vanishes_diff [OF this XY] show "vanishes Y" by simp hoelzl@51523: next hoelzl@51523: assume "vanishes Y" hoelzl@51523: from vanishes_add [OF this XY] show "vanishes X" by simp hoelzl@51523: qed hoelzl@51523: thus "?thesis X Y" hoelzl@51523: unfolding realrel_def hoelzl@51523: by (simp add: vanishes_diff_inverse X Y XY) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "x - y = (x::real) + - y" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "x / y = (x::real) * inverse y" hoelzl@51523: hoelzl@51523: lemma add_Real: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "Real X + Real Y = Real (\n. X n + Y n)" hoelzl@51523: using assms plus_real.transfer hoelzl@51523: unfolding cr_real_eq fun_rel_def by simp hoelzl@51523: hoelzl@51523: lemma minus_Real: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: shows "- Real X = Real (\n. - X n)" hoelzl@51523: using assms uminus_real.transfer hoelzl@51523: unfolding cr_real_eq fun_rel_def by simp hoelzl@51523: hoelzl@51523: lemma diff_Real: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "Real X - Real Y = Real (\n. X n - Y n)" hoelzl@51523: unfolding minus_real_def diff_minus hoelzl@51523: by (simp add: minus_Real add_Real X Y) hoelzl@51523: hoelzl@51523: lemma mult_Real: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "Real X * Real Y = Real (\n. X n * Y n)" hoelzl@51523: using assms times_real.transfer hoelzl@51523: unfolding cr_real_eq fun_rel_def by simp hoelzl@51523: hoelzl@51523: lemma inverse_Real: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: shows "inverse (Real X) = hoelzl@51523: (if vanishes X then 0 else Real (\n. inverse (X n)))" hoelzl@51523: using assms inverse_real.transfer zero_real.transfer hoelzl@51523: unfolding cr_real_eq fun_rel_def by (simp split: split_if_asm, metis) hoelzl@51523: hoelzl@51523: instance proof hoelzl@51523: fix a b c :: real hoelzl@51523: show "a + b = b + a" hoelzl@51523: by transfer (simp add: add_ac realrel_def) hoelzl@51523: show "(a + b) + c = a + (b + c)" hoelzl@51523: by transfer (simp add: add_ac realrel_def) hoelzl@51523: show "0 + a = a" hoelzl@51523: by transfer (simp add: realrel_def) hoelzl@51523: show "- a + a = 0" hoelzl@51523: by transfer (simp add: realrel_def) hoelzl@51523: show "a - b = a + - b" hoelzl@51523: by (rule minus_real_def) hoelzl@51523: show "(a * b) * c = a * (b * c)" hoelzl@51523: by transfer (simp add: mult_ac realrel_def) hoelzl@51523: show "a * b = b * a" hoelzl@51523: by transfer (simp add: mult_ac realrel_def) hoelzl@51523: show "1 * a = a" hoelzl@51523: by transfer (simp add: mult_ac realrel_def) hoelzl@51523: show "(a + b) * c = a * c + b * c" hoelzl@51523: by transfer (simp add: distrib_right realrel_def) hoelzl@51523: show "(0\real) \ (1\real)" hoelzl@51523: by transfer (simp add: realrel_def) hoelzl@51523: show "a \ 0 \ inverse a * a = 1" hoelzl@51523: apply transfer hoelzl@51523: apply (simp add: realrel_def) hoelzl@51523: apply (rule vanishesI) hoelzl@51523: apply (frule (1) cauchy_not_vanishes, clarify) hoelzl@51523: apply (rule_tac x=k in exI, clarify) hoelzl@51523: apply (drule_tac x=n in spec, simp) hoelzl@51523: done hoelzl@51523: show "a / b = a * inverse b" hoelzl@51523: by (rule divide_real_def) hoelzl@51523: show "inverse (0::real) = 0" hoelzl@51523: by transfer (simp add: realrel_def) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: subsection {* Positive reals *} hoelzl@51523: hoelzl@51523: lift_definition positive :: "real \ bool" hoelzl@51523: is "\X. \r>0. \k. \n\k. r < X n" hoelzl@51523: proof - hoelzl@51523: { fix X Y hoelzl@51523: assume "realrel X Y" hoelzl@51523: hence XY: "vanishes (\n. X n - Y n)" hoelzl@51523: unfolding realrel_def by simp_all hoelzl@51523: assume "\r>0. \k. \n\k. r < X n" hoelzl@51523: then obtain r i where "0 < r" and i: "\n\i. r < X n" hoelzl@51523: by fast hoelzl@51523: obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" hoelzl@51523: using `0 < r` by (rule obtain_pos_sum) hoelzl@51523: obtain j where j: "\n\j. \X n - Y n\ < s" hoelzl@51523: using vanishesD [OF XY s] .. hoelzl@51523: have "\n\max i j. t < Y n" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix n assume n: "i \ n" "j \ n" hoelzl@51523: have "\X n - Y n\ < s" and "r < X n" hoelzl@51523: using i j n by simp_all hoelzl@51523: thus "t < Y n" unfolding r by simp hoelzl@51523: qed hoelzl@51523: hence "\r>0. \k. \n\k. r < Y n" using t by fast hoelzl@51523: } note 1 = this hoelzl@51523: fix X Y assume "realrel X Y" hoelzl@51523: hence "realrel X Y" and "realrel Y X" hoelzl@51523: using symp_realrel unfolding symp_def by auto hoelzl@51523: thus "?thesis X Y" hoelzl@51523: by (safe elim!: 1) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma positive_Real: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: shows "positive (Real X) \ (\r>0. \k. \n\k. r < X n)" hoelzl@51523: using assms positive.transfer hoelzl@51523: unfolding cr_real_eq fun_rel_def by simp hoelzl@51523: hoelzl@51523: lemma positive_zero: "\ positive 0" hoelzl@51523: by transfer auto hoelzl@51523: hoelzl@51523: lemma positive_add: hoelzl@51523: "positive x \ positive y \ positive (x + y)" hoelzl@51523: apply transfer hoelzl@51523: apply (clarify, rename_tac a b i j) hoelzl@51523: apply (rule_tac x="a + b" in exI, simp) hoelzl@51523: apply (rule_tac x="max i j" in exI, clarsimp) hoelzl@51523: apply (simp add: add_strict_mono) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma positive_mult: hoelzl@51523: "positive x \ positive y \ positive (x * y)" hoelzl@51523: apply transfer hoelzl@51523: apply (clarify, rename_tac a b i j) hoelzl@51523: apply (rule_tac x="a * b" in exI, simp add: mult_pos_pos) hoelzl@51523: apply (rule_tac x="max i j" in exI, clarsimp) hoelzl@51523: apply (rule mult_strict_mono, auto) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma positive_minus: hoelzl@51523: "\ positive x \ x \ 0 \ positive (- x)" hoelzl@51523: apply transfer hoelzl@51523: apply (simp add: realrel_def) hoelzl@51523: apply (drule (1) cauchy_not_vanishes_cases, safe, fast, fast) hoelzl@51523: done hoelzl@51523: hoelzl@51523: instantiation real :: linordered_field_inverse_zero hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "x < y \ positive (y - x)" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "x \ (y::real) \ x < y \ x = y" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "abs (a::real) = (if a < 0 then - a else a)" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)" hoelzl@51523: hoelzl@51523: instance proof hoelzl@51523: fix a b c :: real hoelzl@51523: show "\a\ = (if a < 0 then - a else a)" hoelzl@51523: by (rule abs_real_def) hoelzl@51523: show "a < b \ a \ b \ \ b \ a" hoelzl@51523: unfolding less_eq_real_def less_real_def hoelzl@51523: by (auto, drule (1) positive_add, simp_all add: positive_zero) hoelzl@51523: show "a \ a" hoelzl@51523: unfolding less_eq_real_def by simp hoelzl@51523: show "a \ b \ b \ c \ a \ c" hoelzl@51523: unfolding less_eq_real_def less_real_def hoelzl@51523: by (auto, drule (1) positive_add, simp add: algebra_simps) hoelzl@51523: show "a \ b \ b \ a \ a = b" hoelzl@51523: unfolding less_eq_real_def less_real_def hoelzl@51523: by (auto, drule (1) positive_add, simp add: positive_zero) hoelzl@51523: show "a \ b \ c + a \ c + b" hoelzl@51523: unfolding less_eq_real_def less_real_def by (auto simp: diff_minus) (* by auto *) hoelzl@51523: (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \ b + - a *) hoelzl@51523: (* Should produce c + b - (c + a) \ b - a *) hoelzl@51523: show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" hoelzl@51523: by (rule sgn_real_def) hoelzl@51523: show "a \ b \ b \ a" hoelzl@51523: unfolding less_eq_real_def less_real_def hoelzl@51523: by (auto dest!: positive_minus) hoelzl@51523: show "a < b \ 0 < c \ c * a < c * b" hoelzl@51523: unfolding less_real_def hoelzl@51523: by (drule (1) positive_mult, simp add: algebra_simps) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: instantiation real :: distrib_lattice hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "(inf :: real \ real \ real) = min" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "(sup :: real \ real \ real) = max" hoelzl@51523: hoelzl@51523: instance proof hoelzl@51523: qed (auto simp add: inf_real_def sup_real_def min_max.sup_inf_distrib1) hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: lemma of_nat_Real: "of_nat x = Real (\n. of_nat x)" hoelzl@51523: apply (induct x) hoelzl@51523: apply (simp add: zero_real_def) hoelzl@51523: apply (simp add: one_real_def add_Real) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma of_int_Real: "of_int x = Real (\n. of_int x)" hoelzl@51523: apply (cases x rule: int_diff_cases) hoelzl@51523: apply (simp add: of_nat_Real diff_Real) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma of_rat_Real: "of_rat x = Real (\n. x)" hoelzl@51523: apply (induct x) hoelzl@51523: apply (simp add: Fract_of_int_quotient of_rat_divide) hoelzl@51523: apply (simp add: of_int_Real divide_inverse) hoelzl@51523: apply (simp add: inverse_Real mult_Real) hoelzl@51523: done hoelzl@51523: hoelzl@51523: instance real :: archimedean_field hoelzl@51523: proof hoelzl@51523: fix x :: real hoelzl@51523: show "\z. x \ of_int z" hoelzl@51523: apply (induct x) hoelzl@51523: apply (frule cauchy_imp_bounded, clarify) hoelzl@51523: apply (rule_tac x="ceiling b + 1" in exI) hoelzl@51523: apply (rule less_imp_le) hoelzl@51523: apply (simp add: of_int_Real less_real_def diff_Real positive_Real) hoelzl@51523: apply (rule_tac x=1 in exI, simp add: algebra_simps) hoelzl@51523: apply (rule_tac x=0 in exI, clarsimp) hoelzl@51523: apply (rule le_less_trans [OF abs_ge_self]) hoelzl@51523: apply (rule less_le_trans [OF _ le_of_int_ceiling]) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: qed hoelzl@51523: hoelzl@51523: instantiation real :: floor_ceiling hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition [code del]: hoelzl@51523: "floor (x::real) = (THE z. of_int z \ x \ x < of_int (z + 1))" hoelzl@51523: hoelzl@51523: instance proof hoelzl@51523: fix x :: real hoelzl@51523: show "of_int (floor x) \ x \ x < of_int (floor x + 1)" hoelzl@51523: unfolding floor_real_def using floor_exists1 by (rule theI') hoelzl@51523: qed hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: subsection {* Completeness *} hoelzl@51523: hoelzl@51523: lemma not_positive_Real: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: shows "\ positive (Real X) \ (\r>0. \k. \n\k. X n \ r)" hoelzl@51523: unfolding positive_Real [OF X] hoelzl@51523: apply (auto, unfold not_less) hoelzl@51523: apply (erule obtain_pos_sum) hoelzl@51523: apply (drule_tac x=s in spec, simp) hoelzl@51523: apply (drule_tac r=t in cauchyD [OF X], clarify) hoelzl@51523: apply (drule_tac x=k in spec, clarsimp) hoelzl@51523: apply (rule_tac x=n in exI, clarify, rename_tac m) hoelzl@51523: apply (drule_tac x=m in spec, simp) hoelzl@51523: apply (drule_tac x=n in spec, simp) hoelzl@51523: apply (drule spec, drule (1) mp, clarify, rename_tac i) hoelzl@51523: apply (rule_tac x="max i k" in exI, simp) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma le_Real: hoelzl@51523: assumes X: "cauchy X" and Y: "cauchy Y" hoelzl@51523: shows "Real X \ Real Y = (\r>0. \k. \n\k. X n \ Y n + r)" hoelzl@51523: unfolding not_less [symmetric, where 'a=real] less_real_def hoelzl@51523: apply (simp add: diff_Real not_positive_Real X Y) hoelzl@51523: apply (simp add: diff_le_eq add_ac) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma le_RealI: hoelzl@51523: assumes Y: "cauchy Y" hoelzl@51523: shows "\n. x \ of_rat (Y n) \ x \ Real Y" hoelzl@51523: proof (induct x) hoelzl@51523: fix X assume X: "cauchy X" and "\n. Real X \ of_rat (Y n)" hoelzl@51523: hence le: "\m r. 0 < r \ \k. \n\k. X n \ Y m + r" hoelzl@51523: by (simp add: of_rat_Real le_Real) hoelzl@51523: { hoelzl@51523: fix r :: rat assume "0 < r" hoelzl@51523: then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" hoelzl@51523: by (rule obtain_pos_sum) hoelzl@51523: obtain i where i: "\m\i. \n\i. \Y m - Y n\ < s" hoelzl@51523: using cauchyD [OF Y s] .. hoelzl@51523: obtain j where j: "\n\j. X n \ Y i + t" hoelzl@51523: using le [OF t] .. hoelzl@51523: have "\n\max i j. X n \ Y n + r" hoelzl@51523: proof (clarsimp) hoelzl@51523: fix n assume n: "i \ n" "j \ n" hoelzl@51523: have "X n \ Y i + t" using n j by simp hoelzl@51523: moreover have "\Y i - Y n\ < s" using n i by simp hoelzl@51523: ultimately show "X n \ Y n + r" unfolding r by simp hoelzl@51523: qed hoelzl@51523: hence "\k. \n\k. X n \ Y n + r" .. hoelzl@51523: } hoelzl@51523: thus "Real X \ Real Y" hoelzl@51523: by (simp add: of_rat_Real le_Real X Y) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma Real_leI: hoelzl@51523: assumes X: "cauchy X" hoelzl@51523: assumes le: "\n. of_rat (X n) \ y" hoelzl@51523: shows "Real X \ y" hoelzl@51523: proof - hoelzl@51523: have "- y \ - Real X" hoelzl@51523: by (simp add: minus_Real X le_RealI of_rat_minus le) hoelzl@51523: thus ?thesis by simp hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma less_RealD: hoelzl@51523: assumes Y: "cauchy Y" hoelzl@51523: shows "x < Real Y \ \n. x < of_rat (Y n)" hoelzl@51523: by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y]) hoelzl@51523: hoelzl@51523: lemma of_nat_less_two_power: hoelzl@51523: "of_nat n < (2::'a::linordered_idom) ^ n" hoelzl@51523: apply (induct n) hoelzl@51523: apply simp hoelzl@51523: apply (subgoal_tac "(1::'a) \ 2 ^ n") hoelzl@51523: apply (drule (1) add_le_less_mono, simp) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma complete_real: hoelzl@51523: fixes S :: "real set" hoelzl@51523: assumes "\x. x \ S" and "\z. \x\S. x \ z" hoelzl@51523: shows "\y. (\x\S. x \ y) \ (\z. (\x\S. x \ z) \ y \ z)" hoelzl@51523: proof - hoelzl@51523: obtain x where x: "x \ S" using assms(1) .. hoelzl@51523: obtain z where z: "\x\S. x \ z" using assms(2) .. hoelzl@51523: hoelzl@51523: def P \ "\x. \y\S. y \ of_rat x" hoelzl@51523: obtain a where a: "\ P a" hoelzl@51523: proof hoelzl@51523: have "of_int (floor (x - 1)) \ x - 1" by (rule of_int_floor_le) hoelzl@51523: also have "x - 1 < x" by simp hoelzl@51523: finally have "of_int (floor (x - 1)) < x" . hoelzl@51523: hence "\ x \ of_int (floor (x - 1))" by (simp only: not_le) hoelzl@51523: then show "\ P (of_int (floor (x - 1)))" hoelzl@51523: unfolding P_def of_rat_of_int_eq using x by fast hoelzl@51523: qed hoelzl@51523: obtain b where b: "P b" hoelzl@51523: proof hoelzl@51523: show "P (of_int (ceiling z))" hoelzl@51523: unfolding P_def of_rat_of_int_eq hoelzl@51523: proof hoelzl@51523: fix y assume "y \ S" hoelzl@51523: hence "y \ z" using z by simp hoelzl@51523: also have "z \ of_int (ceiling z)" by (rule le_of_int_ceiling) hoelzl@51523: finally show "y \ of_int (ceiling z)" . hoelzl@51523: qed hoelzl@51523: qed hoelzl@51523: hoelzl@51523: def avg \ "\x y :: rat. x/2 + y/2" hoelzl@51523: def bisect \ "\(x, y). if P (avg x y) then (x, avg x y) else (avg x y, y)" hoelzl@51523: def A \ "\n. fst ((bisect ^^ n) (a, b))" hoelzl@51523: def B \ "\n. snd ((bisect ^^ n) (a, b))" hoelzl@51523: def C \ "\n. avg (A n) (B n)" hoelzl@51523: have A_0 [simp]: "A 0 = a" unfolding A_def by simp hoelzl@51523: have B_0 [simp]: "B 0 = b" unfolding B_def by simp hoelzl@51523: have A_Suc [simp]: "\n. A (Suc n) = (if P (C n) then A n else C n)" hoelzl@51523: unfolding A_def B_def C_def bisect_def split_def by simp hoelzl@51523: have B_Suc [simp]: "\n. B (Suc n) = (if P (C n) then C n else B n)" hoelzl@51523: unfolding A_def B_def C_def bisect_def split_def by simp hoelzl@51523: hoelzl@51523: have width: "\n. B n - A n = (b - a) / 2^n" hoelzl@51523: apply (simp add: eq_divide_eq) hoelzl@51523: apply (induct_tac n, simp) hoelzl@51523: apply (simp add: C_def avg_def algebra_simps) hoelzl@51523: done hoelzl@51523: hoelzl@51523: have twos: "\y r :: rat. 0 < r \ \n. y / 2 ^ n < r" hoelzl@51523: apply (simp add: divide_less_eq) hoelzl@51523: apply (subst mult_commute) hoelzl@51523: apply (frule_tac y=y in ex_less_of_nat_mult) hoelzl@51523: apply clarify hoelzl@51523: apply (rule_tac x=n in exI) hoelzl@51523: apply (erule less_trans) hoelzl@51523: apply (rule mult_strict_right_mono) hoelzl@51523: apply (rule le_less_trans [OF _ of_nat_less_two_power]) hoelzl@51523: apply simp hoelzl@51523: apply assumption hoelzl@51523: done hoelzl@51523: hoelzl@51523: have PA: "\n. \ P (A n)" hoelzl@51523: by (induct_tac n, simp_all add: a) hoelzl@51523: have PB: "\n. P (B n)" hoelzl@51523: by (induct_tac n, simp_all add: b) hoelzl@51523: have ab: "a < b" hoelzl@51523: using a b unfolding P_def hoelzl@51523: apply (clarsimp simp add: not_le) hoelzl@51523: apply (drule (1) bspec) hoelzl@51523: apply (drule (1) less_le_trans) hoelzl@51523: apply (simp add: of_rat_less) hoelzl@51523: done hoelzl@51523: have AB: "\n. A n < B n" hoelzl@51523: by (induct_tac n, simp add: ab, simp add: C_def avg_def) hoelzl@51523: have A_mono: "\i j. i \ j \ A i \ A j" hoelzl@51523: apply (auto simp add: le_less [where 'a=nat]) hoelzl@51523: apply (erule less_Suc_induct) hoelzl@51523: apply (clarsimp simp add: C_def avg_def) hoelzl@51523: apply (simp add: add_divide_distrib [symmetric]) hoelzl@51523: apply (rule AB [THEN less_imp_le]) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: have B_mono: "\i j. i \ j \ B j \ B i" hoelzl@51523: apply (auto simp add: le_less [where 'a=nat]) hoelzl@51523: apply (erule less_Suc_induct) hoelzl@51523: apply (clarsimp simp add: C_def avg_def) hoelzl@51523: apply (simp add: add_divide_distrib [symmetric]) hoelzl@51523: apply (rule AB [THEN less_imp_le]) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: have cauchy_lemma: hoelzl@51523: "\X. \n. \i\n. A n \ X i \ X i \ B n \ cauchy X" hoelzl@51523: apply (rule cauchyI) hoelzl@51523: apply (drule twos [where y="b - a"]) hoelzl@51523: apply (erule exE) hoelzl@51523: apply (rule_tac x=n in exI, clarify, rename_tac i j) hoelzl@51523: apply (rule_tac y="B n - A n" in le_less_trans) defer hoelzl@51523: apply (simp add: width) hoelzl@51523: apply (drule_tac x=n in spec) hoelzl@51523: apply (frule_tac x=i in spec, drule (1) mp) hoelzl@51523: apply (frule_tac x=j in spec, drule (1) mp) hoelzl@51523: apply (frule A_mono, drule B_mono) hoelzl@51523: apply (frule A_mono, drule B_mono) hoelzl@51523: apply arith hoelzl@51523: done hoelzl@51523: have "cauchy A" hoelzl@51523: apply (rule cauchy_lemma [rule_format]) hoelzl@51523: apply (simp add: A_mono) hoelzl@51523: apply (erule order_trans [OF less_imp_le [OF AB] B_mono]) hoelzl@51523: done hoelzl@51523: have "cauchy B" hoelzl@51523: apply (rule cauchy_lemma [rule_format]) hoelzl@51523: apply (simp add: B_mono) hoelzl@51523: apply (erule order_trans [OF A_mono less_imp_le [OF AB]]) hoelzl@51523: done hoelzl@51523: have 1: "\x\S. x \ Real B" hoelzl@51523: proof hoelzl@51523: fix x assume "x \ S" hoelzl@51523: then show "x \ Real B" hoelzl@51523: using PB [unfolded P_def] `cauchy B` hoelzl@51523: by (simp add: le_RealI) hoelzl@51523: qed hoelzl@51523: have 2: "\z. (\x\S. x \ z) \ Real A \ z" hoelzl@51523: apply clarify hoelzl@51523: apply (erule contrapos_pp) hoelzl@51523: apply (simp add: not_le) hoelzl@51523: apply (drule less_RealD [OF `cauchy A`], clarify) hoelzl@51523: apply (subgoal_tac "\ P (A n)") hoelzl@51523: apply (simp add: P_def not_le, clarify) hoelzl@51523: apply (erule rev_bexI) hoelzl@51523: apply (erule (1) less_trans) hoelzl@51523: apply (simp add: PA) hoelzl@51523: done hoelzl@51523: have "vanishes (\n. (b - a) / 2 ^ n)" hoelzl@51523: proof (rule vanishesI) hoelzl@51523: fix r :: rat assume "0 < r" hoelzl@51523: then obtain k where k: "\b - a\ / 2 ^ k < r" hoelzl@51523: using twos by fast hoelzl@51523: have "\n\k. \(b - a) / 2 ^ n\ < r" hoelzl@51523: proof (clarify) hoelzl@51523: fix n assume n: "k \ n" hoelzl@51523: have "\(b - a) / 2 ^ n\ = \b - a\ / 2 ^ n" hoelzl@51523: by simp hoelzl@51523: also have "\ \ \b - a\ / 2 ^ k" hoelzl@51523: using n by (simp add: divide_left_mono mult_pos_pos) hoelzl@51523: also note k hoelzl@51523: finally show "\(b - a) / 2 ^ n\ < r" . hoelzl@51523: qed hoelzl@51523: thus "\k. \n\k. \(b - a) / 2 ^ n\ < r" .. hoelzl@51523: qed hoelzl@51523: hence 3: "Real B = Real A" hoelzl@51523: by (simp add: eq_Real `cauchy A` `cauchy B` width) hoelzl@51523: show "\y. (\x\S. x \ y) \ (\z. (\x\S. x \ z) \ y \ z)" hoelzl@51523: using 1 2 3 by (rule_tac x="Real B" in exI, simp) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: hoelzl@51773: instantiation real :: conditionally_complete_linorder hoelzl@51523: begin hoelzl@51523: hoelzl@51523: subsection{*Supremum of a set of reals*} hoelzl@51523: hoelzl@51523: definition hoelzl@51523: Sup_real_def: "Sup X \ LEAST z::real. \x\X. x\z" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: Inf_real_def: "Inf (X::real set) \ - Sup (uminus ` X)" hoelzl@51523: hoelzl@51523: instance hoelzl@51523: proof hoelzl@51523: { fix z x :: real and X :: "real set" hoelzl@51523: assume x: "x \ X" and z: "\x. x \ X \ x \ z" hoelzl@51523: then obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" hoelzl@51523: using complete_real[of X] by blast hoelzl@51523: then show "x \ Sup X" hoelzl@51523: unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) } hoelzl@51523: note Sup_upper = this hoelzl@51523: hoelzl@51523: { fix z :: real and X :: "real set" hoelzl@51523: assume x: "X \ {}" and z: "\x. x \ X \ x \ z" hoelzl@51523: then obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" hoelzl@51523: using complete_real[of X] by blast hoelzl@51523: then have "Sup X = s" hoelzl@51523: unfolding Sup_real_def by (best intro: Least_equality) hoelzl@51523: also with s z have "... \ z" hoelzl@51523: by blast hoelzl@51523: finally show "Sup X \ z" . } hoelzl@51523: note Sup_least = this hoelzl@51523: hoelzl@51523: { fix x z :: real and X :: "real set" hoelzl@51523: assume x: "x \ X" and z: "\x. x \ X \ z \ x" hoelzl@51523: have "-x \ Sup (uminus ` X)" hoelzl@51523: by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z) hoelzl@51523: then show "Inf X \ x" hoelzl@51523: by (auto simp add: Inf_real_def) } hoelzl@51523: hoelzl@51523: { fix z :: real and X :: "real set" hoelzl@51523: assume x: "X \ {}" and z: "\x. x \ X \ z \ x" hoelzl@51523: have "Sup (uminus ` X) \ -z" hoelzl@51523: using x z by (force intro: Sup_least) hoelzl@51523: then show "z \ Inf X" hoelzl@51523: by (auto simp add: Inf_real_def) } hoelzl@51523: qed hoelzl@51523: end hoelzl@51523: hoelzl@51523: text {* hoelzl@51523: \medskip Completeness properties using @{text "isUb"}, @{text "isLub"}: hoelzl@51523: *} hoelzl@51523: hoelzl@51523: lemma reals_complete: "\X. X \ S \ \Y. isUb (UNIV::real set) S Y \ \t. isLub (UNIV :: real set) S t" hoelzl@51523: by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro: cSup_upper) hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection {* Hiding implementation details *} hoelzl@51523: hoelzl@51523: hide_const (open) vanishes cauchy positive Real hoelzl@51523: hoelzl@51523: declare Real_induct [induct del] hoelzl@51523: declare Abs_real_induct [induct del] hoelzl@51523: declare Abs_real_cases [cases del] hoelzl@51523: hoelzl@51523: lemmas [transfer_rule del] = hoelzl@51523: real.All_transfer real.Ex_transfer real.rel_eq_transfer forall_real_transfer hoelzl@51523: zero_real.transfer one_real.transfer plus_real.transfer uminus_real.transfer hoelzl@51523: times_real.transfer inverse_real.transfer positive.transfer real.right_unique hoelzl@51523: real.right_total hoelzl@51523: hoelzl@51523: subsection{*More Lemmas*} hoelzl@51523: hoelzl@51523: text {* BH: These lemmas should not be necessary; they should be hoelzl@51523: covered by existing simp rules and simplification procedures. *} hoelzl@51523: hoelzl@51523: lemma real_mult_left_cancel: "(c::real) \ 0 ==> (c*a=c*b) = (a=b)" hoelzl@51523: by simp (* redundant with mult_cancel_left *) hoelzl@51523: hoelzl@51523: lemma real_mult_right_cancel: "(c::real) \ 0 ==> (a*c=b*c) = (a=b)" hoelzl@51523: by simp (* redundant with mult_cancel_right *) hoelzl@51523: hoelzl@51523: lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)" hoelzl@51523: by simp (* solved by linordered_ring_less_cancel_factor simproc *) hoelzl@51523: hoelzl@51523: lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \ y*z) = (x\y)" hoelzl@51523: by simp (* solved by linordered_ring_le_cancel_factor simproc *) hoelzl@51523: hoelzl@51523: lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \ z*y) = (x\y)" hoelzl@51523: by simp (* solved by linordered_ring_le_cancel_factor simproc *) hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection {* Embedding numbers into the Reals *} hoelzl@51523: hoelzl@51523: abbreviation hoelzl@51523: real_of_nat :: "nat \ real" hoelzl@51523: where hoelzl@51523: "real_of_nat \ of_nat" hoelzl@51523: hoelzl@51523: abbreviation hoelzl@51523: real_of_int :: "int \ real" hoelzl@51523: where hoelzl@51523: "real_of_int \ of_int" hoelzl@51523: hoelzl@51523: abbreviation hoelzl@51523: real_of_rat :: "rat \ real" hoelzl@51523: where hoelzl@51523: "real_of_rat \ of_rat" hoelzl@51523: hoelzl@51523: consts hoelzl@51523: (*overloaded constant for injecting other types into "real"*) hoelzl@51523: real :: "'a => real" hoelzl@51523: hoelzl@51523: defs (overloaded) hoelzl@51523: real_of_nat_def [code_unfold]: "real == real_of_nat" hoelzl@51523: real_of_int_def [code_unfold]: "real == real_of_int" hoelzl@51523: hoelzl@51523: declare [[coercion_enabled]] hoelzl@51523: declare [[coercion "real::nat\real"]] hoelzl@51523: declare [[coercion "real::int\real"]] hoelzl@51523: declare [[coercion "int"]] hoelzl@51523: hoelzl@51523: declare [[coercion_map map]] hoelzl@51523: declare [[coercion_map "% f g h x. g (h (f x))"]] hoelzl@51523: declare [[coercion_map "% f g (x,y) . (f x, g y)"]] hoelzl@51523: hoelzl@51523: lemma real_eq_of_nat: "real = of_nat" hoelzl@51523: unfolding real_of_nat_def .. hoelzl@51523: hoelzl@51523: lemma real_eq_of_int: "real = of_int" hoelzl@51523: unfolding real_of_int_def .. hoelzl@51523: hoelzl@51523: lemma real_of_int_zero [simp]: "real (0::int) = 0" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_one [simp]: "real (1::int) = (1::real)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n" hoelzl@51523: by (simp add: real_of_int_def of_int_power) hoelzl@51523: hoelzl@51523: lemmas power_real_of_int = real_of_int_power [symmetric] hoelzl@51523: hoelzl@51523: lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))" hoelzl@51523: apply (subst real_eq_of_int)+ hoelzl@51523: apply (rule of_int_setsum) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) = hoelzl@51523: (PROD x:A. real(f x))" hoelzl@51523: apply (subst real_eq_of_int)+ hoelzl@51523: apply (rule of_int_setprod) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \ real y) = (x \ y)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)" hoelzl@51523: by (simp add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \ 1 < i" hoelzl@51523: unfolding real_of_one[symmetric] real_of_int_less_iff .. hoelzl@51523: hoelzl@51523: lemma one_le_real_of_int_cancel_iff: "1 \ real (i :: int) \ 1 \ i" hoelzl@51523: unfolding real_of_one[symmetric] real_of_int_le_iff .. hoelzl@51523: hoelzl@51523: lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \ i < 1" hoelzl@51523: unfolding real_of_one[symmetric] real_of_int_less_iff .. hoelzl@51523: hoelzl@51523: lemma real_of_int_le_one_cancel_iff: "real (i :: int) \ 1 \ i \ 1" hoelzl@51523: unfolding real_of_one[symmetric] real_of_int_le_iff .. hoelzl@51523: hoelzl@51523: lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))" hoelzl@51523: by (auto simp add: abs_if) hoelzl@51523: hoelzl@51523: lemma int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)" hoelzl@51523: apply (subgoal_tac "real n + 1 = real (n + 1)") hoelzl@51523: apply (simp del: real_of_int_add) hoelzl@51523: apply auto hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)" hoelzl@51523: apply (subgoal_tac "real m + 1 = real (m + 1)") hoelzl@51523: apply (simp del: real_of_int_add) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_div_aux: "(real (x::int)) / (real d) = hoelzl@51523: real (x div d) + (real (x mod d)) / (real d)" hoelzl@51523: proof - hoelzl@51523: have "x = (x div d) * d + x mod d" hoelzl@51523: by auto hoelzl@51523: then have "real x = real (x div d) * real d + real(x mod d)" hoelzl@51523: by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym]) hoelzl@51523: then have "real x / real d = ... / real d" hoelzl@51523: by simp hoelzl@51523: then show ?thesis hoelzl@51523: by (auto simp add: add_divide_distrib algebra_simps) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma real_of_int_div: "(d :: int) dvd n ==> hoelzl@51523: real(n div d) = real n / real d" hoelzl@51523: apply (subst real_of_int_div_aux) hoelzl@51523: apply simp hoelzl@51523: apply (simp add: dvd_eq_mod_eq_0) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_div2: hoelzl@51523: "0 <= real (n::int) / real (x) - real (n div x)" hoelzl@51523: apply (case_tac "x = 0") hoelzl@51523: apply simp hoelzl@51523: apply (case_tac "0 < x") hoelzl@51523: apply (simp add: algebra_simps) hoelzl@51523: apply (subst real_of_int_div_aux) hoelzl@51523: apply simp hoelzl@51523: apply (subst zero_le_divide_iff) hoelzl@51523: apply auto hoelzl@51523: apply (simp add: algebra_simps) hoelzl@51523: apply (subst real_of_int_div_aux) hoelzl@51523: apply simp hoelzl@51523: apply (subst zero_le_divide_iff) hoelzl@51523: apply auto hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_div3: hoelzl@51523: "real (n::int) / real (x) - real (n div x) <= 1" hoelzl@51523: apply (simp add: algebra_simps) hoelzl@51523: apply (subst real_of_int_div_aux) hoelzl@51523: apply (auto simp add: divide_le_eq intro: order_less_imp_le) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" hoelzl@51523: by (insert real_of_int_div2 [of n x], simp) hoelzl@51523: hoelzl@51523: lemma Ints_real_of_int [simp]: "real (x::int) \ Ints" hoelzl@51523: unfolding real_of_int_def by (rule Ints_of_int) hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection{*Embedding the Naturals into the Reals*} hoelzl@51523: hoelzl@51523: lemma real_of_nat_zero [simp]: "real (0::nat) = 0" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_1 [simp]: "real (1::nat) = 1" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: (*Not for addsimps: often the LHS is used to represent a positive natural*) hoelzl@51523: lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_less_iff [iff]: hoelzl@51523: "(real (n::nat) < real m) = (n < m)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_le_iff [iff]: "(real (n::nat) \ real m) = (n \ m)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_ge_zero [iff]: "0 \ real (n::nat)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)" hoelzl@51523: by (simp add: real_of_nat_def del: of_nat_Suc) hoelzl@51523: hoelzl@51523: lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n" hoelzl@51523: by (simp add: real_of_nat_def of_nat_mult) hoelzl@51523: hoelzl@51523: lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n" hoelzl@51523: by (simp add: real_of_nat_def of_nat_power) hoelzl@51523: hoelzl@51523: lemmas power_real_of_nat = real_of_nat_power [symmetric] hoelzl@51523: hoelzl@51523: lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) = hoelzl@51523: (SUM x:A. real(f x))" hoelzl@51523: apply (subst real_eq_of_nat)+ hoelzl@51523: apply (rule of_nat_setsum) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) = hoelzl@51523: (PROD x:A. real(f x))" hoelzl@51523: apply (subst real_eq_of_nat)+ hoelzl@51523: apply (rule of_nat_setprod) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_card: "real (card A) = setsum (%x.1) A" hoelzl@51523: apply (subst card_eq_setsum) hoelzl@51523: apply (subst real_of_nat_setsum) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_diff: "n \ m ==> real (m - n) = real (m::nat) - real n" hoelzl@51523: by (simp add: add: real_of_nat_def of_nat_diff) hoelzl@51523: hoelzl@51523: lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)" hoelzl@51523: by (auto simp: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \ 0) = (n = 0)" hoelzl@51523: by (simp add: add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0" hoelzl@51523: by (simp add: add: real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)" hoelzl@51523: apply (subgoal_tac "real n + 1 = real (Suc n)") hoelzl@51523: apply simp hoelzl@51523: apply (auto simp add: real_of_nat_Suc) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)" hoelzl@51523: apply (subgoal_tac "real m + 1 = real (Suc m)") hoelzl@51523: apply (simp add: less_Suc_eq_le) hoelzl@51523: apply (simp add: real_of_nat_Suc) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_nat_div_aux: "(real (x::nat)) / (real d) = hoelzl@51523: real (x div d) + (real (x mod d)) / (real d)" hoelzl@51523: proof - hoelzl@51523: have "x = (x div d) * d + x mod d" hoelzl@51523: by auto hoelzl@51523: then have "real x = real (x div d) * real d + real(x mod d)" hoelzl@51523: by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym]) hoelzl@51523: then have "real x / real d = \ / real d" hoelzl@51523: by simp hoelzl@51523: then show ?thesis hoelzl@51523: by (auto simp add: add_divide_distrib algebra_simps) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma real_of_nat_div: "(d :: nat) dvd n ==> hoelzl@51523: real(n div d) = real n / real d" hoelzl@51523: by (subst real_of_nat_div_aux) hoelzl@51523: (auto simp add: dvd_eq_mod_eq_0 [symmetric]) hoelzl@51523: hoelzl@51523: lemma real_of_nat_div2: hoelzl@51523: "0 <= real (n::nat) / real (x) - real (n div x)" hoelzl@51523: apply (simp add: algebra_simps) hoelzl@51523: apply (subst real_of_nat_div_aux) hoelzl@51523: apply simp hoelzl@51523: apply (subst zero_le_divide_iff) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_nat_div3: hoelzl@51523: "real (n::nat) / real (x) - real (n div x) <= 1" hoelzl@51523: apply(case_tac "x = 0") hoelzl@51523: apply (simp) hoelzl@51523: apply (simp add: algebra_simps) hoelzl@51523: apply (subst real_of_nat_div_aux) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" hoelzl@51523: by (insert real_of_nat_div2 [of n x], simp) hoelzl@51523: hoelzl@51523: lemma real_of_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n" hoelzl@51523: by (simp add: real_of_int_def real_of_nat_def) hoelzl@51523: hoelzl@51523: lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x" hoelzl@51523: apply (subgoal_tac "real(int(nat x)) = real(nat x)") hoelzl@51523: apply force hoelzl@51523: apply (simp only: real_of_int_of_nat_eq) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma Nats_real_of_nat [simp]: "real (n::nat) \ Nats" hoelzl@51523: unfolding real_of_nat_def by (rule of_nat_in_Nats) hoelzl@51523: hoelzl@51523: lemma Ints_real_of_nat [simp]: "real (n::nat) \ Ints" hoelzl@51523: unfolding real_of_nat_def by (rule Ints_of_nat) hoelzl@51523: hoelzl@51523: subsection {* The Archimedean Property of the Reals *} hoelzl@51523: hoelzl@51523: theorem reals_Archimedean: hoelzl@51523: assumes x_pos: "0 < x" hoelzl@51523: shows "\n. inverse (real (Suc n)) < x" hoelzl@51523: unfolding real_of_nat_def using x_pos hoelzl@51523: by (rule ex_inverse_of_nat_Suc_less) hoelzl@51523: hoelzl@51523: lemma reals_Archimedean2: "\n. (x::real) < real (n::nat)" hoelzl@51523: unfolding real_of_nat_def by (rule ex_less_of_nat) hoelzl@51523: hoelzl@51523: lemma reals_Archimedean3: hoelzl@51523: assumes x_greater_zero: "0 < x" hoelzl@51523: shows "\(y::real). \(n::nat). y < real n * x" hoelzl@51523: unfolding real_of_nat_def using `0 < x` hoelzl@51523: by (auto intro: ex_less_of_nat_mult) hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection{* Rationals *} hoelzl@51523: hoelzl@51523: lemma Rats_real_nat[simp]: "real(n::nat) \ \" hoelzl@51523: by (simp add: real_eq_of_nat) hoelzl@51523: hoelzl@51523: hoelzl@51523: lemma Rats_eq_int_div_int: hoelzl@51523: "\ = { real(i::int)/real(j::int) |i j. j \ 0}" (is "_ = ?S") hoelzl@51523: proof hoelzl@51523: show "\ \ ?S" hoelzl@51523: proof hoelzl@51523: fix x::real assume "x : \" hoelzl@51523: then obtain r where "x = of_rat r" unfolding Rats_def .. hoelzl@51523: have "of_rat r : ?S" hoelzl@51523: by (cases r)(auto simp add:of_rat_rat real_eq_of_int) hoelzl@51523: thus "x : ?S" using `x = of_rat r` by simp hoelzl@51523: qed hoelzl@51523: next hoelzl@51523: show "?S \ \" hoelzl@51523: proof(auto simp:Rats_def) hoelzl@51523: fix i j :: int assume "j \ 0" hoelzl@51523: hence "real i / real j = of_rat(Fract i j)" hoelzl@51523: by (simp add:of_rat_rat real_eq_of_int) hoelzl@51523: thus "real i / real j \ range of_rat" by blast hoelzl@51523: qed hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma Rats_eq_int_div_nat: hoelzl@51523: "\ = { real(i::int)/real(n::nat) |i n. n \ 0}" hoelzl@51523: proof(auto simp:Rats_eq_int_div_int) hoelzl@51523: fix i j::int assume "j \ 0" hoelzl@51523: show "EX (i'::int) (n::nat). real i/real j = real i'/real n \ 00" hoelzl@51523: hence "real i/real j = real i/real(nat j) \ 00" hoelzl@51523: hence "real i/real j = real(-i)/real(nat(-j)) \ 00` hoelzl@51523: by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat) hoelzl@51523: thus ?thesis by blast hoelzl@51523: qed hoelzl@51523: next hoelzl@51523: fix i::int and n::nat assume "0 < n" hoelzl@51523: hence "real i/real n = real i/real(int n) \ int n \ 0" by simp hoelzl@51523: thus "\(i'::int) j::int. real i/real n = real i'/real j \ j \ 0" by blast hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma Rats_abs_nat_div_natE: hoelzl@51523: assumes "x \ \" hoelzl@51523: obtains m n :: nat hoelzl@51523: where "n \ 0" and "\x\ = real m / real n" and "gcd m n = 1" hoelzl@51523: proof - hoelzl@51523: from `x \ \` obtain i::int and n::nat where "n \ 0" and "x = real i / real n" hoelzl@51523: by(auto simp add: Rats_eq_int_div_nat) hoelzl@51523: hence "\x\ = real(nat(abs i)) / real n" by simp hoelzl@51523: then obtain m :: nat where x_rat: "\x\ = real m / real n" by blast hoelzl@51523: let ?gcd = "gcd m n" hoelzl@51523: from `n\0` have gcd: "?gcd \ 0" by simp hoelzl@51523: let ?k = "m div ?gcd" hoelzl@51523: let ?l = "n div ?gcd" hoelzl@51523: let ?gcd' = "gcd ?k ?l" hoelzl@51523: have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m" hoelzl@51523: by (rule dvd_mult_div_cancel) hoelzl@51523: have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n" hoelzl@51523: by (rule dvd_mult_div_cancel) hoelzl@51523: from `n\0` and gcd_l have "?l \ 0" by (auto iff del: neq0_conv) hoelzl@51523: moreover hoelzl@51523: have "\x\ = real ?k / real ?l" hoelzl@51523: proof - hoelzl@51523: from gcd have "real ?k / real ?l = hoelzl@51523: real (?gcd * ?k) / real (?gcd * ?l)" by simp hoelzl@51523: also from gcd_k and gcd_l have "\ = real m / real n" by simp hoelzl@51523: also from x_rat have "\ = \x\" .. hoelzl@51523: finally show ?thesis .. hoelzl@51523: qed hoelzl@51523: moreover hoelzl@51523: have "?gcd' = 1" hoelzl@51523: proof - hoelzl@51523: have "?gcd * ?gcd' = gcd (?gcd * ?k) (?gcd * ?l)" hoelzl@51523: by (rule gcd_mult_distrib_nat) hoelzl@51523: with gcd_k gcd_l have "?gcd * ?gcd' = ?gcd" by simp hoelzl@51523: with gcd show ?thesis by auto hoelzl@51523: qed hoelzl@51523: ultimately show ?thesis .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: subsection{*Density of the Rational Reals in the Reals*} hoelzl@51523: hoelzl@51523: text{* This density proof is due to Stefan Richter and was ported by TN. The hoelzl@51523: original source is \emph{Real Analysis} by H.L. Royden. hoelzl@51523: It employs the Archimedean property of the reals. *} hoelzl@51523: hoelzl@51523: lemma Rats_dense_in_real: hoelzl@51523: fixes x :: real hoelzl@51523: assumes "x < y" shows "\r\\. x < r \ r < y" hoelzl@51523: proof - hoelzl@51523: from `x "ceiling (y * real q) - 1" hoelzl@51523: def r \ "of_int p / real q" hoelzl@51523: from q have "x < y - inverse (real q)" by simp hoelzl@51523: also have "y - inverse (real q) \ r" hoelzl@51523: unfolding r_def p_def hoelzl@51523: by (simp add: le_divide_eq left_diff_distrib le_of_int_ceiling `0 < q`) hoelzl@51523: finally have "x < r" . hoelzl@51523: moreover have "r < y" hoelzl@51523: unfolding r_def p_def hoelzl@51523: by (simp add: divide_less_eq diff_less_eq `0 < q` hoelzl@51523: less_ceiling_iff [symmetric]) hoelzl@51523: moreover from r_def have "r \ \" by simp hoelzl@51523: ultimately show ?thesis by fast hoelzl@51523: qed hoelzl@51523: hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection{*Numerals and Arithmetic*} hoelzl@51523: hoelzl@51523: lemma [code_abbrev]: hoelzl@51523: "real_of_int (numeral k) = numeral k" hoelzl@51523: "real_of_int (neg_numeral k) = neg_numeral k" hoelzl@51523: by simp_all hoelzl@51523: hoelzl@51523: text{*Collapse applications of @{term real} to @{term number_of}*} hoelzl@51523: lemma real_numeral [simp]: hoelzl@51523: "real (numeral v :: int) = numeral v" hoelzl@51523: "real (neg_numeral v :: int) = neg_numeral v" hoelzl@51523: by (simp_all add: real_of_int_def) hoelzl@51523: hoelzl@51523: lemma real_of_nat_numeral [simp]: hoelzl@51523: "real (numeral v :: nat) = numeral v" hoelzl@51523: by (simp add: real_of_nat_def) hoelzl@51523: hoelzl@51523: declaration {* hoelzl@51523: K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2] hoelzl@51523: (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *) hoelzl@51523: #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2] hoelzl@51523: (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *) hoelzl@51523: #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add}, hoelzl@51523: @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one}, hoelzl@51523: @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff}, hoelzl@51523: @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq}, hoelzl@51523: @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}] hoelzl@51523: #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \ real"}) hoelzl@51523: #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \ real"})) hoelzl@51523: *} hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*} hoelzl@51523: hoelzl@51523: lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" hoelzl@51523: by arith hoelzl@51523: hoelzl@51523: text {* FIXME: redundant with @{text add_eq_0_iff} below *} hoelzl@51523: lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: lemma real_add_le_0_iff: "(x+y \ (0::real)) = (y \ -x)" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: lemma real_0_le_add_iff: "((0::real) \ x+y) = (-x \ y)" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: subsection {* Lemmas about powers *} hoelzl@51523: hoelzl@51523: text {* FIXME: declare this in Rings.thy or not at all *} hoelzl@51523: declare abs_mult_self [simp] hoelzl@51523: hoelzl@51523: (* used by Import/HOL/real.imp *) hoelzl@51523: lemma two_realpow_ge_one: "(1::real) \ 2 ^ n" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n" hoelzl@51523: apply (induct "n") hoelzl@51523: apply (auto simp add: real_of_nat_Suc) hoelzl@51523: apply (subst mult_2) hoelzl@51523: apply (erule add_less_le_mono) hoelzl@51523: apply (rule two_realpow_ge_one) hoelzl@51523: done hoelzl@51523: hoelzl@51523: text {* TODO: no longer real-specific; rename and move elsewhere *} hoelzl@51523: lemma realpow_Suc_le_self: hoelzl@51523: fixes r :: "'a::linordered_semidom" hoelzl@51523: shows "[| 0 \ r; r \ 1 |] ==> r ^ Suc n \ r" hoelzl@51523: by (insert power_decreasing [of 1 "Suc n" r], simp) hoelzl@51523: hoelzl@51523: text {* TODO: no longer real-specific; rename and move elsewhere *} hoelzl@51523: lemma realpow_minus_mult: hoelzl@51523: fixes x :: "'a::monoid_mult" hoelzl@51523: shows "0 < n \ x ^ (n - 1) * x = x ^ n" hoelzl@51523: by (simp add: power_commutes split add: nat_diff_split) hoelzl@51523: hoelzl@51523: text {* FIXME: declare this [simp] for all types, or not at all *} hoelzl@51523: lemma real_two_squares_add_zero_iff [simp]: hoelzl@51523: "(x * x + y * y = 0) = ((x::real) = 0 \ y = 0)" hoelzl@51523: by (rule sum_squares_eq_zero_iff) hoelzl@51523: hoelzl@51523: text {* FIXME: declare this [simp] for all types, or not at all *} hoelzl@51523: lemma realpow_two_sum_zero_iff [simp]: hoelzl@51523: "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)" hoelzl@51523: by (rule sum_power2_eq_zero_iff) hoelzl@51523: hoelzl@51523: lemma real_minus_mult_self_le [simp]: "-(u * u) \ (x * (x::real))" hoelzl@51523: by (rule_tac y = 0 in order_trans, auto) hoelzl@51523: hoelzl@51523: lemma realpow_square_minus_le [simp]: "-(u ^ 2) \ (x::real) ^ 2" hoelzl@51523: by (auto simp add: power2_eq_square) hoelzl@51523: hoelzl@51523: hoelzl@51523: lemma numeral_power_le_real_of_nat_cancel_iff[simp]: hoelzl@51523: "(numeral x::real) ^ n \ real a \ (numeral x::nat) ^ n \ a" hoelzl@51523: unfolding real_of_nat_le_iff[symmetric] by simp hoelzl@51523: hoelzl@51523: lemma real_of_nat_le_numeral_power_cancel_iff[simp]: hoelzl@51523: "real a \ (numeral x::real) ^ n \ a \ (numeral x::nat) ^ n" hoelzl@51523: unfolding real_of_nat_le_iff[symmetric] by simp hoelzl@51523: hoelzl@51523: lemma numeral_power_le_real_of_int_cancel_iff[simp]: hoelzl@51523: "(numeral x::real) ^ n \ real a \ (numeral x::int) ^ n \ a" hoelzl@51523: unfolding real_of_int_le_iff[symmetric] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_le_numeral_power_cancel_iff[simp]: hoelzl@51523: "real a \ (numeral x::real) ^ n \ a \ (numeral x::int) ^ n" hoelzl@51523: unfolding real_of_int_le_iff[symmetric] by simp hoelzl@51523: hoelzl@51523: lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]: hoelzl@51523: "(neg_numeral x::real) ^ n \ real a \ (neg_numeral x::int) ^ n \ a" hoelzl@51523: unfolding real_of_int_le_iff[symmetric] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]: hoelzl@51523: "real a \ (neg_numeral x::real) ^ n \ a \ (neg_numeral x::int) ^ n" hoelzl@51523: unfolding real_of_int_le_iff[symmetric] by simp hoelzl@51523: hoelzl@51523: subsection{*Density of the Reals*} hoelzl@51523: hoelzl@51523: lemma real_lbound_gt_zero: hoelzl@51523: "[| (0::real) < d1; 0 < d2 |] ==> \e. 0 < e & e < d1 & e < d2" hoelzl@51523: apply (rule_tac x = " (min d1 d2) /2" in exI) hoelzl@51523: apply (simp add: min_def) hoelzl@51523: done hoelzl@51523: hoelzl@51523: hoelzl@51523: text{*Similar results are proved in @{text Fields}*} hoelzl@51523: lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y" hoelzl@51523: by auto hoelzl@51523: hoelzl@51523: lemma real_sum_of_halves: "x/2 + x/2 = (x::real)" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: subsection{*Absolute Value Function for the Reals*} hoelzl@51523: hoelzl@51523: lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))" hoelzl@51523: by (simp add: abs_if) hoelzl@51523: hoelzl@51523: (* FIXME: redundant, but used by Integration/RealRandVar.thy in AFP *) hoelzl@51523: lemma abs_le_interval_iff: "(abs x \ r) = (-r\x & x\(r::real))" hoelzl@51523: by (force simp add: abs_le_iff) hoelzl@51523: hoelzl@51523: lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)" hoelzl@51523: by (simp add: abs_if) hoelzl@51523: hoelzl@51523: lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)" hoelzl@51523: by (rule abs_of_nonneg [OF real_of_nat_ge_zero]) hoelzl@51523: hoelzl@51523: lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \ abs(x + -l) + abs(y + -m)" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection{*Floor and Ceiling Functions from the Reals to the Integers*} hoelzl@51523: hoelzl@51523: (* FIXME: theorems for negative numerals *) hoelzl@51523: lemma numeral_less_real_of_int_iff [simp]: hoelzl@51523: "((numeral n) < real (m::int)) = (numeral n < m)" hoelzl@51523: apply auto hoelzl@51523: apply (rule real_of_int_less_iff [THEN iffD1]) hoelzl@51523: apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma numeral_less_real_of_int_iff2 [simp]: hoelzl@51523: "(real (m::int) < (numeral n)) = (m < numeral n)" hoelzl@51523: apply auto hoelzl@51523: apply (rule real_of_int_less_iff [THEN iffD1]) hoelzl@51523: apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma numeral_le_real_of_int_iff [simp]: hoelzl@51523: "((numeral n) \ real (m::int)) = (numeral n \ m)" hoelzl@51523: by (simp add: linorder_not_less [symmetric]) hoelzl@51523: hoelzl@51523: lemma numeral_le_real_of_int_iff2 [simp]: hoelzl@51523: "(real (m::int) \ (numeral n)) = (m \ numeral n)" hoelzl@51523: by (simp add: linorder_not_less [symmetric]) hoelzl@51523: hoelzl@51523: lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n" hoelzl@51523: unfolding real_of_nat_def by simp hoelzl@51523: hoelzl@51523: lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n" hoelzl@51523: unfolding real_of_nat_def by (simp add: floor_minus) hoelzl@51523: hoelzl@51523: lemma floor_real_of_int [simp]: "floor (real (n::int)) = n" hoelzl@51523: unfolding real_of_int_def by simp hoelzl@51523: hoelzl@51523: lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n" hoelzl@51523: unfolding real_of_int_def by (simp add: floor_minus) hoelzl@51523: hoelzl@51523: lemma real_lb_ub_int: " \n::int. real n \ r & r < real (n+1)" hoelzl@51523: unfolding real_of_int_def by (rule floor_exists) hoelzl@51523: hoelzl@51523: lemma lemma_floor: hoelzl@51523: assumes a1: "real m \ r" and a2: "r < real n + 1" hoelzl@51523: shows "m \ (n::int)" hoelzl@51523: proof - hoelzl@51523: have "real m < real n + 1" using a1 a2 by (rule order_le_less_trans) hoelzl@51523: also have "... = real (n + 1)" by simp hoelzl@51523: finally have "m < n + 1" by (simp only: real_of_int_less_iff) hoelzl@51523: thus ?thesis by arith hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma real_of_int_floor_le [simp]: "real (floor r) \ r" hoelzl@51523: unfolding real_of_int_def by (rule of_int_floor_le) hoelzl@51523: hoelzl@51523: lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \ x" hoelzl@51523: by (auto intro: lemma_floor) hoelzl@51523: hoelzl@51523: lemma real_of_int_floor_cancel [simp]: hoelzl@51523: "(real (floor x) = x) = (\n::int. x = real n)" hoelzl@51523: using floor_real_of_int by metis hoelzl@51523: hoelzl@51523: lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n" hoelzl@51523: unfolding real_of_int_def using floor_unique [of n x] by simp hoelzl@51523: hoelzl@51523: lemma floor_eq2: "[| real n \ x; x < real n + 1 |] ==> floor x = n" hoelzl@51523: unfolding real_of_int_def by (rule floor_unique) hoelzl@51523: hoelzl@51523: lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n" hoelzl@51523: apply (rule inj_int [THEN injD]) hoelzl@51523: apply (simp add: real_of_nat_Suc) hoelzl@51523: apply (simp add: real_of_nat_Suc floor_eq floor_eq [where n = "int n"]) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma floor_eq4: "[| real n \ x; x < real (Suc n) |] ==> nat(floor x) = n" hoelzl@51523: apply (drule order_le_imp_less_or_eq) hoelzl@51523: apply (auto intro: floor_eq3) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real(floor r)" hoelzl@51523: unfolding real_of_int_def using floor_correct [of r] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)" hoelzl@51523: unfolding real_of_int_def using floor_correct [of r] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_floor_add_one_ge [simp]: "r \ real(floor r) + 1" hoelzl@51523: unfolding real_of_int_def using floor_correct [of r] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1" hoelzl@51523: unfolding real_of_int_def using floor_correct [of r] by simp hoelzl@51523: hoelzl@51523: lemma le_floor: "real a <= x ==> a <= floor x" hoelzl@51523: unfolding real_of_int_def by (simp add: le_floor_iff) hoelzl@51523: hoelzl@51523: lemma real_le_floor: "a <= floor x ==> real a <= x" hoelzl@51523: unfolding real_of_int_def by (simp add: le_floor_iff) hoelzl@51523: hoelzl@51523: lemma le_floor_eq: "(a <= floor x) = (real a <= x)" hoelzl@51523: unfolding real_of_int_def by (rule le_floor_iff) hoelzl@51523: hoelzl@51523: lemma floor_less_eq: "(floor x < a) = (x < real a)" hoelzl@51523: unfolding real_of_int_def by (rule floor_less_iff) hoelzl@51523: hoelzl@51523: lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)" hoelzl@51523: unfolding real_of_int_def by (rule less_floor_iff) hoelzl@51523: hoelzl@51523: lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)" hoelzl@51523: unfolding real_of_int_def by (rule floor_le_iff) hoelzl@51523: hoelzl@51523: lemma floor_add [simp]: "floor (x + real a) = floor x + a" hoelzl@51523: unfolding real_of_int_def by (rule floor_add_of_int) hoelzl@51523: hoelzl@51523: lemma floor_subtract [simp]: "floor (x - real a) = floor x - a" hoelzl@51523: unfolding real_of_int_def by (rule floor_diff_of_int) hoelzl@51523: hoelzl@51523: lemma le_mult_floor: hoelzl@51523: assumes "0 \ (a :: real)" and "0 \ b" hoelzl@51523: shows "floor a * floor b \ floor (a * b)" hoelzl@51523: proof - hoelzl@51523: have "real (floor a) \ a" hoelzl@51523: and "real (floor b) \ b" by auto hoelzl@51523: hence "real (floor a * floor b) \ a * b" hoelzl@51523: using assms by (auto intro!: mult_mono) hoelzl@51523: also have "a * b < real (floor (a * b) + 1)" by auto hoelzl@51523: finally show ?thesis unfolding real_of_int_less_iff by simp hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma floor_divide_eq_div: hoelzl@51523: "floor (real a / real b) = a div b" hoelzl@51523: proof cases hoelzl@51523: assume "b \ 0 \ b dvd a" hoelzl@51523: with real_of_int_div3[of a b] show ?thesis hoelzl@51523: by (auto simp: real_of_int_div[symmetric] intro!: floor_eq2 real_of_int_div4 neq_le_trans) hoelzl@51523: (metis add_left_cancel zero_neq_one real_of_int_div_aux real_of_int_inject hoelzl@51523: real_of_int_zero_cancel right_inverse_eq div_self mod_div_trivial) hoelzl@51523: qed (auto simp: real_of_int_div) hoelzl@51523: hoelzl@51523: lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n" hoelzl@51523: unfolding real_of_nat_def by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_ceiling_ge [simp]: "r \ real (ceiling r)" hoelzl@51523: unfolding real_of_int_def by (rule le_of_int_ceiling) hoelzl@51523: hoelzl@51523: lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n" hoelzl@51523: unfolding real_of_int_def by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_ceiling_cancel [simp]: hoelzl@51523: "(real (ceiling x) = x) = (\n::int. x = real n)" hoelzl@51523: using ceiling_real_of_int by metis hoelzl@51523: hoelzl@51523: lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1" hoelzl@51523: unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp hoelzl@51523: hoelzl@51523: lemma ceiling_eq2: "[| real n < x; x \ real n + 1 |] ==> ceiling x = n + 1" hoelzl@51523: unfolding real_of_int_def using ceiling_unique [of "n + 1" x] by simp hoelzl@51523: hoelzl@51523: lemma ceiling_eq3: "[| real n - 1 < x; x \ real n |] ==> ceiling x = n" hoelzl@51523: unfolding real_of_int_def using ceiling_unique [of n x] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \ r" hoelzl@51523: unfolding real_of_int_def using ceiling_correct [of r] by simp hoelzl@51523: hoelzl@51523: lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \ r + 1" hoelzl@51523: unfolding real_of_int_def using ceiling_correct [of r] by simp hoelzl@51523: hoelzl@51523: lemma ceiling_le: "x <= real a ==> ceiling x <= a" hoelzl@51523: unfolding real_of_int_def by (simp add: ceiling_le_iff) hoelzl@51523: hoelzl@51523: lemma ceiling_le_real: "ceiling x <= a ==> x <= real a" hoelzl@51523: unfolding real_of_int_def by (simp add: ceiling_le_iff) hoelzl@51523: hoelzl@51523: lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)" hoelzl@51523: unfolding real_of_int_def by (rule ceiling_le_iff) hoelzl@51523: hoelzl@51523: lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)" hoelzl@51523: unfolding real_of_int_def by (rule less_ceiling_iff) hoelzl@51523: hoelzl@51523: lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)" hoelzl@51523: unfolding real_of_int_def by (rule ceiling_less_iff) hoelzl@51523: hoelzl@51523: lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)" hoelzl@51523: unfolding real_of_int_def by (rule le_ceiling_iff) hoelzl@51523: hoelzl@51523: lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a" hoelzl@51523: unfolding real_of_int_def by (rule ceiling_add_of_int) hoelzl@51523: hoelzl@51523: lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a" hoelzl@51523: unfolding real_of_int_def by (rule ceiling_diff_of_int) hoelzl@51523: hoelzl@51523: hoelzl@51523: subsubsection {* Versions for the natural numbers *} hoelzl@51523: hoelzl@51523: definition hoelzl@51523: natfloor :: "real => nat" where hoelzl@51523: "natfloor x = nat(floor x)" hoelzl@51523: hoelzl@51523: definition hoelzl@51523: natceiling :: "real => nat" where hoelzl@51523: "natceiling x = nat(ceiling x)" hoelzl@51523: hoelzl@51523: lemma natfloor_zero [simp]: "natfloor 0 = 0" hoelzl@51523: by (unfold natfloor_def, simp) hoelzl@51523: hoelzl@51523: lemma natfloor_one [simp]: "natfloor 1 = 1" hoelzl@51523: by (unfold natfloor_def, simp) hoelzl@51523: hoelzl@51523: lemma zero_le_natfloor [simp]: "0 <= natfloor x" hoelzl@51523: by (unfold natfloor_def, simp) hoelzl@51523: hoelzl@51523: lemma natfloor_numeral_eq [simp]: "natfloor (numeral n) = numeral n" hoelzl@51523: by (unfold natfloor_def, simp) hoelzl@51523: hoelzl@51523: lemma natfloor_real_of_nat [simp]: "natfloor(real n) = n" hoelzl@51523: by (unfold natfloor_def, simp) hoelzl@51523: hoelzl@51523: lemma real_natfloor_le: "0 <= x ==> real(natfloor x) <= x" hoelzl@51523: by (unfold natfloor_def, simp) hoelzl@51523: hoelzl@51523: lemma natfloor_neg: "x <= 0 ==> natfloor x = 0" hoelzl@51523: unfolding natfloor_def by simp hoelzl@51523: hoelzl@51523: lemma natfloor_mono: "x <= y ==> natfloor x <= natfloor y" hoelzl@51523: unfolding natfloor_def by (intro nat_mono floor_mono) hoelzl@51523: hoelzl@51523: lemma le_natfloor: "real x <= a ==> x <= natfloor a" hoelzl@51523: apply (unfold natfloor_def) hoelzl@51523: apply (subst nat_int [THEN sym]) hoelzl@51523: apply (rule nat_mono) hoelzl@51523: apply (rule le_floor) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma natfloor_less_iff: "0 \ x \ natfloor x < n \ x < real n" hoelzl@51523: unfolding natfloor_def real_of_nat_def hoelzl@51523: by (simp add: nat_less_iff floor_less_iff) hoelzl@51523: hoelzl@51523: lemma less_natfloor: hoelzl@51523: assumes "0 \ x" and "x < real (n :: nat)" hoelzl@51523: shows "natfloor x < n" hoelzl@51523: using assms by (simp add: natfloor_less_iff) hoelzl@51523: hoelzl@51523: lemma le_natfloor_eq: "0 <= x ==> (a <= natfloor x) = (real a <= x)" hoelzl@51523: apply (rule iffI) hoelzl@51523: apply (rule order_trans) hoelzl@51523: prefer 2 hoelzl@51523: apply (erule real_natfloor_le) hoelzl@51523: apply (subst real_of_nat_le_iff) hoelzl@51523: apply assumption hoelzl@51523: apply (erule le_natfloor) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma le_natfloor_eq_numeral [simp]: hoelzl@51523: "~ neg((numeral n)::int) ==> 0 <= x ==> hoelzl@51523: (numeral n <= natfloor x) = (numeral n <= x)" hoelzl@51523: apply (subst le_natfloor_eq, assumption) hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma le_natfloor_eq_one [simp]: "(1 <= natfloor x) = (1 <= x)" hoelzl@51523: apply (case_tac "0 <= x") hoelzl@51523: apply (subst le_natfloor_eq, assumption, simp) hoelzl@51523: apply (rule iffI) hoelzl@51523: apply (subgoal_tac "natfloor x <= natfloor 0") hoelzl@51523: apply simp hoelzl@51523: apply (rule natfloor_mono) hoelzl@51523: apply simp hoelzl@51523: apply simp hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma natfloor_eq: "real n <= x ==> x < real n + 1 ==> natfloor x = n" hoelzl@51523: unfolding natfloor_def by (simp add: floor_eq2 [where n="int n"]) hoelzl@51523: hoelzl@51523: lemma real_natfloor_add_one_gt: "x < real(natfloor x) + 1" hoelzl@51523: apply (case_tac "0 <= x") hoelzl@51523: apply (unfold natfloor_def) hoelzl@51523: apply simp hoelzl@51523: apply simp_all hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma real_natfloor_gt_diff_one: "x - 1 < real(natfloor x)" hoelzl@51523: using real_natfloor_add_one_gt by (simp add: algebra_simps) hoelzl@51523: hoelzl@51523: lemma ge_natfloor_plus_one_imp_gt: "natfloor z + 1 <= n ==> z < real n" hoelzl@51523: apply (subgoal_tac "z < real(natfloor z) + 1") hoelzl@51523: apply arith hoelzl@51523: apply (rule real_natfloor_add_one_gt) hoelzl@51523: done hoelzl@51523: hoelzl@51523: lemma natfloor_add [simp]: "0 <= x ==> natfloor (x + real a) = natfloor x + a" hoelzl@51523: unfolding natfloor_def hoelzl@51523: unfolding real_of_int_of_nat_eq [symmetric] floor_add hoelzl@51523: by (simp add: nat_add_distrib) hoelzl@51523: hoelzl@51523: lemma natfloor_add_numeral [simp]: hoelzl@51523: "~neg ((numeral n)::int) ==> 0 <= x ==> hoelzl@51523: natfloor (x + numeral n) = natfloor x + numeral n" hoelzl@51523: by (simp add: natfloor_add [symmetric]) hoelzl@51523: hoelzl@51523: lemma natfloor_add_one: "0 <= x ==> natfloor(x + 1) = natfloor x + 1" hoelzl@51523: by (simp add: natfloor_add [symmetric] del: One_nat_def) hoelzl@51523: hoelzl@51523: lemma natfloor_subtract [simp]: hoelzl@51523: "natfloor(x - real a) = natfloor x - a" hoelzl@51523: unfolding natfloor_def real_of_int_of_nat_eq [symmetric] floor_subtract hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma natfloor_div_nat: hoelzl@51523: assumes "1 <= x" and "y > 0" hoelzl@51523: shows "natfloor (x / real y) = natfloor x div y" hoelzl@51523: proof (rule natfloor_eq) hoelzl@51523: have "(natfloor x) div y * y \ natfloor x" hoelzl@51523: by (rule add_leD1 [where k="natfloor x mod y"], simp) hoelzl@51523: thus "real (natfloor x div y) \ x / real y" hoelzl@51523: using assms by (simp add: le_divide_eq le_natfloor_eq) hoelzl@51523: have "natfloor x < (natfloor x) div y * y + y" hoelzl@51523: apply (subst mod_div_equality [symmetric]) hoelzl@51523: apply (rule add_strict_left_mono) hoelzl@51523: apply (rule mod_less_divisor) hoelzl@51523: apply fact hoelzl@51523: done hoelzl@51523: thus "x / real y < real (natfloor x div y) + 1" hoelzl@51523: using assms hoelzl@51523: by (simp add: divide_less_eq natfloor_less_iff distrib_right) hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma le_mult_natfloor: hoelzl@51523: shows "natfloor a * natfloor b \ natfloor (a * b)" hoelzl@51523: by (cases "0 <= a & 0 <= b") hoelzl@51523: (auto simp add: le_natfloor_eq mult_nonneg_nonneg mult_mono' real_natfloor_le natfloor_neg) hoelzl@51523: hoelzl@51523: lemma natceiling_zero [simp]: "natceiling 0 = 0" hoelzl@51523: by (unfold natceiling_def, simp) hoelzl@51523: hoelzl@51523: lemma natceiling_one [simp]: "natceiling 1 = 1" hoelzl@51523: by (unfold natceiling_def, simp) hoelzl@51523: hoelzl@51523: lemma zero_le_natceiling [simp]: "0 <= natceiling x" hoelzl@51523: by (unfold natceiling_def, simp) hoelzl@51523: hoelzl@51523: lemma natceiling_numeral_eq [simp]: "natceiling (numeral n) = numeral n" hoelzl@51523: by (unfold natceiling_def, simp) hoelzl@51523: hoelzl@51523: lemma natceiling_real_of_nat [simp]: "natceiling(real n) = n" hoelzl@51523: by (unfold natceiling_def, simp) hoelzl@51523: hoelzl@51523: lemma real_natceiling_ge: "x <= real(natceiling x)" hoelzl@51523: unfolding natceiling_def by (cases "x < 0", simp_all) hoelzl@51523: hoelzl@51523: lemma natceiling_neg: "x <= 0 ==> natceiling x = 0" hoelzl@51523: unfolding natceiling_def by simp hoelzl@51523: hoelzl@51523: lemma natceiling_mono: "x <= y ==> natceiling x <= natceiling y" hoelzl@51523: unfolding natceiling_def by (intro nat_mono ceiling_mono) hoelzl@51523: hoelzl@51523: lemma natceiling_le: "x <= real a ==> natceiling x <= a" hoelzl@51523: unfolding natceiling_def real_of_nat_def hoelzl@51523: by (simp add: nat_le_iff ceiling_le_iff) hoelzl@51523: hoelzl@51523: lemma natceiling_le_eq: "(natceiling x <= a) = (x <= real a)" hoelzl@51523: unfolding natceiling_def real_of_nat_def hoelzl@51523: by (simp add: nat_le_iff ceiling_le_iff) hoelzl@51523: hoelzl@51523: lemma natceiling_le_eq_numeral [simp]: hoelzl@51523: "~ neg((numeral n)::int) ==> hoelzl@51523: (natceiling x <= numeral n) = (x <= numeral n)" hoelzl@51523: by (simp add: natceiling_le_eq) hoelzl@51523: hoelzl@51523: lemma natceiling_le_eq_one: "(natceiling x <= 1) = (x <= 1)" hoelzl@51523: unfolding natceiling_def hoelzl@51523: by (simp add: nat_le_iff ceiling_le_iff) hoelzl@51523: hoelzl@51523: lemma natceiling_eq: "real n < x ==> x <= real n + 1 ==> natceiling x = n + 1" hoelzl@51523: unfolding natceiling_def hoelzl@51523: by (simp add: ceiling_eq2 [where n="int n"]) hoelzl@51523: hoelzl@51523: lemma natceiling_add [simp]: "0 <= x ==> hoelzl@51523: natceiling (x + real a) = natceiling x + a" hoelzl@51523: unfolding natceiling_def hoelzl@51523: unfolding real_of_int_of_nat_eq [symmetric] ceiling_add hoelzl@51523: by (simp add: nat_add_distrib) hoelzl@51523: hoelzl@51523: lemma natceiling_add_numeral [simp]: hoelzl@51523: "~ neg ((numeral n)::int) ==> 0 <= x ==> hoelzl@51523: natceiling (x + numeral n) = natceiling x + numeral n" hoelzl@51523: by (simp add: natceiling_add [symmetric]) hoelzl@51523: hoelzl@51523: lemma natceiling_add_one: "0 <= x ==> natceiling(x + 1) = natceiling x + 1" hoelzl@51523: by (simp add: natceiling_add [symmetric] del: One_nat_def) hoelzl@51523: hoelzl@51523: lemma natceiling_subtract [simp]: "natceiling(x - real a) = natceiling x - a" hoelzl@51523: unfolding natceiling_def real_of_int_of_nat_eq [symmetric] ceiling_subtract hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: subsection {* Exponentiation with floor *} hoelzl@51523: hoelzl@51523: lemma floor_power: hoelzl@51523: assumes "x = real (floor x)" hoelzl@51523: shows "floor (x ^ n) = floor x ^ n" hoelzl@51523: proof - hoelzl@51523: have *: "x ^ n = real (floor x ^ n)" hoelzl@51523: using assms by (induct n arbitrary: x) simp_all hoelzl@51523: show ?thesis unfolding real_of_int_inject[symmetric] hoelzl@51523: unfolding * floor_real_of_int .. hoelzl@51523: qed hoelzl@51523: hoelzl@51523: lemma natfloor_power: hoelzl@51523: assumes "x = real (natfloor x)" hoelzl@51523: shows "natfloor (x ^ n) = natfloor x ^ n" hoelzl@51523: proof - hoelzl@51523: from assms have "0 \ floor x" by auto hoelzl@51523: note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \ floor x`]] hoelzl@51523: from floor_power[OF this] hoelzl@51523: show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \ floor x`, symmetric] hoelzl@51523: by simp hoelzl@51523: qed hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection {* Implementation of rational real numbers *} hoelzl@51523: hoelzl@51523: text {* Formal constructor *} hoelzl@51523: hoelzl@51523: definition Ratreal :: "rat \ real" where hoelzl@51523: [code_abbrev, simp]: "Ratreal = of_rat" hoelzl@51523: hoelzl@51523: code_datatype Ratreal hoelzl@51523: hoelzl@51523: hoelzl@51523: text {* Numerals *} hoelzl@51523: hoelzl@51523: lemma [code_abbrev]: hoelzl@51523: "(of_rat (of_int a) :: real) = of_int a" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma [code_abbrev]: hoelzl@51523: "(of_rat 0 :: real) = 0" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma [code_abbrev]: hoelzl@51523: "(of_rat 1 :: real) = 1" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma [code_abbrev]: hoelzl@51523: "(of_rat (numeral k) :: real) = numeral k" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma [code_abbrev]: hoelzl@51523: "(of_rat (neg_numeral k) :: real) = neg_numeral k" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma [code_post]: hoelzl@51523: "(of_rat (0 / r) :: real) = 0" hoelzl@51523: "(of_rat (r / 0) :: real) = 0" hoelzl@51523: "(of_rat (1 / 1) :: real) = 1" hoelzl@51523: "(of_rat (numeral k / 1) :: real) = numeral k" hoelzl@51523: "(of_rat (neg_numeral k / 1) :: real) = neg_numeral k" hoelzl@51523: "(of_rat (1 / numeral k) :: real) = 1 / numeral k" hoelzl@51523: "(of_rat (1 / neg_numeral k) :: real) = 1 / neg_numeral k" hoelzl@51523: "(of_rat (numeral k / numeral l) :: real) = numeral k / numeral l" hoelzl@51523: "(of_rat (numeral k / neg_numeral l) :: real) = numeral k / neg_numeral l" hoelzl@51523: "(of_rat (neg_numeral k / numeral l) :: real) = neg_numeral k / numeral l" hoelzl@51523: "(of_rat (neg_numeral k / neg_numeral l) :: real) = neg_numeral k / neg_numeral l" hoelzl@51523: by (simp_all add: of_rat_divide) hoelzl@51523: hoelzl@51523: hoelzl@51523: text {* Operations *} hoelzl@51523: hoelzl@51523: lemma zero_real_code [code]: hoelzl@51523: "0 = Ratreal 0" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: lemma one_real_code [code]: hoelzl@51523: "1 = Ratreal 1" hoelzl@51523: by simp hoelzl@51523: hoelzl@51523: instantiation real :: equal hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition "HOL.equal (x\real) y \ x - y = 0" hoelzl@51523: hoelzl@51523: instance proof hoelzl@51523: qed (simp add: equal_real_def) hoelzl@51523: hoelzl@51523: lemma real_equal_code [code]: hoelzl@51523: "HOL.equal (Ratreal x) (Ratreal y) \ HOL.equal x y" hoelzl@51523: by (simp add: equal_real_def equal) hoelzl@51523: hoelzl@51523: lemma [code nbe]: hoelzl@51523: "HOL.equal (x::real) x \ True" hoelzl@51523: by (rule equal_refl) hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: lemma real_less_eq_code [code]: "Ratreal x \ Ratreal y \ x \ y" hoelzl@51523: by (simp add: of_rat_less_eq) hoelzl@51523: hoelzl@51523: lemma real_less_code [code]: "Ratreal x < Ratreal y \ x < y" hoelzl@51523: by (simp add: of_rat_less) hoelzl@51523: hoelzl@51523: lemma real_plus_code [code]: "Ratreal x + Ratreal y = Ratreal (x + y)" hoelzl@51523: by (simp add: of_rat_add) hoelzl@51523: hoelzl@51523: lemma real_times_code [code]: "Ratreal x * Ratreal y = Ratreal (x * y)" hoelzl@51523: by (simp add: of_rat_mult) hoelzl@51523: hoelzl@51523: lemma real_uminus_code [code]: "- Ratreal x = Ratreal (- x)" hoelzl@51523: by (simp add: of_rat_minus) hoelzl@51523: hoelzl@51523: lemma real_minus_code [code]: "Ratreal x - Ratreal y = Ratreal (x - y)" hoelzl@51523: by (simp add: of_rat_diff) hoelzl@51523: hoelzl@51523: lemma real_inverse_code [code]: "inverse (Ratreal x) = Ratreal (inverse x)" hoelzl@51523: by (simp add: of_rat_inverse) hoelzl@51523: hoelzl@51523: lemma real_divide_code [code]: "Ratreal x / Ratreal y = Ratreal (x / y)" hoelzl@51523: by (simp add: of_rat_divide) hoelzl@51523: hoelzl@51523: lemma real_floor_code [code]: "floor (Ratreal x) = floor x" hoelzl@51523: by (metis Ratreal_def floor_le_iff floor_unique le_floor_iff of_int_floor_le of_rat_of_int_eq real_less_eq_code) hoelzl@51523: hoelzl@51523: hoelzl@51523: text {* Quickcheck *} hoelzl@51523: hoelzl@51523: definition (in term_syntax) hoelzl@51523: valterm_ratreal :: "rat \ (unit \ Code_Evaluation.term) \ real \ (unit \ Code_Evaluation.term)" where hoelzl@51523: [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\} k" hoelzl@51523: hoelzl@51523: notation fcomp (infixl "\>" 60) hoelzl@51523: notation scomp (infixl "\\" 60) hoelzl@51523: hoelzl@51523: instantiation real :: random hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\r. Pair (valterm_ratreal r))" hoelzl@51523: hoelzl@51523: instance .. hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: no_notation fcomp (infixl "\>" 60) hoelzl@51523: no_notation scomp (infixl "\\" 60) hoelzl@51523: hoelzl@51523: instantiation real :: exhaustive hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d" hoelzl@51523: hoelzl@51523: instance .. hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: instantiation real :: full_exhaustive hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d" hoelzl@51523: hoelzl@51523: instance .. hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: instantiation real :: narrowing hoelzl@51523: begin hoelzl@51523: hoelzl@51523: definition hoelzl@51523: "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing" hoelzl@51523: hoelzl@51523: instance .. hoelzl@51523: hoelzl@51523: end hoelzl@51523: hoelzl@51523: hoelzl@51523: subsection {* Setup for Nitpick *} hoelzl@51523: hoelzl@51523: declaration {* hoelzl@51523: Nitpick_HOL.register_frac_type @{type_name real} hoelzl@51523: [(@{const_name zero_real_inst.zero_real}, @{const_name Nitpick.zero_frac}), hoelzl@51523: (@{const_name one_real_inst.one_real}, @{const_name Nitpick.one_frac}), hoelzl@51523: (@{const_name plus_real_inst.plus_real}, @{const_name Nitpick.plus_frac}), hoelzl@51523: (@{const_name times_real_inst.times_real}, @{const_name Nitpick.times_frac}), hoelzl@51523: (@{const_name uminus_real_inst.uminus_real}, @{const_name Nitpick.uminus_frac}), hoelzl@51523: (@{const_name inverse_real_inst.inverse_real}, @{const_name Nitpick.inverse_frac}), hoelzl@51523: (@{const_name ord_real_inst.less_real}, @{const_name Nitpick.less_frac}), hoelzl@51523: (@{const_name ord_real_inst.less_eq_real}, @{const_name Nitpick.less_eq_frac})] hoelzl@51523: *} hoelzl@51523: hoelzl@51523: lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real hoelzl@51523: ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real hoelzl@51523: times_real_inst.times_real uminus_real_inst.uminus_real hoelzl@51523: zero_real_inst.zero_real hoelzl@51523: hoelzl@51523: ML_file "Tools/SMT/smt_real.ML" hoelzl@51523: setup SMT_Real.setup hoelzl@51523: hoelzl@51523: end