# HG changeset patch # User wenzelm # Date 1466716117 -7200 # Node ID 176d1f408696eec6c1d473dfcea9d99dad4cc7d5 # Parent 4eaf35781b23b9cd8893626f2ba41e3ed5cf9bd8 misc tuning and modernization; diff -r 4eaf35781b23 -r 176d1f408696 src/HOL/Real.thy --- a/src/HOL/Real.thy Thu Jun 23 11:01:14 2016 +0200 +++ b/src/HOL/Real.thy Thu Jun 23 23:08:37 2016 +0200 @@ -20,28 +20,22 @@ construction using Dedekind cuts. \ + subsection \Preliminary lemmas\ -lemma inj_add_left [simp]: - fixes x :: "'a::cancel_semigroup_add" shows "inj (op+ x)" -by (meson add_left_imp_eq injI) +lemma inj_add_left [simp]: "inj (op + x)" for x :: "'a::cancel_semigroup_add" + by (meson add_left_imp_eq injI) -lemma inj_mult_left [simp]: "inj (op* x) \ x \ (0::'a::idom)" +lemma inj_mult_left [simp]: "inj (op * x) \ x \ 0" for x :: "'a::idom" by (metis injI mult_cancel_left the_inv_f_f zero_neq_one) -lemma add_diff_add: - fixes a b c d :: "'a::ab_group_add" - shows "(a + c) - (b + d) = (a - b) + (c - d)" +lemma add_diff_add: "(a + c) - (b + d) = (a - b) + (c - d)" for a b c d :: "'a::ab_group_add" by simp -lemma minus_diff_minus: - fixes a b :: "'a::ab_group_add" - shows "- a - - b = - (a - b)" +lemma minus_diff_minus: "- a - - b = - (a - b)" for a b :: "'a::ab_group_add" by simp -lemma mult_diff_mult: - fixes x y a b :: "'a::ring" - shows "(x * y - a * b) = x * (y - b) + (x - a) * b" +lemma mult_diff_mult: "(x * y - a * b) = x * (y - b) + (x - a) * b" for x y a b :: "'a::ring" by (simp add: algebra_simps) lemma inverse_diff_inverse: @@ -54,38 +48,41 @@ fixes r :: rat assumes r: "0 < r" obtains s t where "0 < s" and "0 < t" and "r = s + t" proof - from r show "0 < r/2" by simp - from r show "0 < r/2" by simp - show "r = r/2 + r/2" by simp + from r show "0 < r/2" by simp + from r show "0 < r/2" by simp + show "r = r/2 + r/2" by simp qed + subsection \Sequences that converge to zero\ -definition - vanishes :: "(nat \ rat) \ bool" -where - "vanishes X = (\r>0. \k. \n\k. \X n\ < r)" +definition vanishes :: "(nat \ rat) \ bool" + where "vanishes X \ (\r>0. \k. \n\k. \X n\ < r)" lemma vanishesI: "(\r. 0 < r \ \k. \n\k. \X n\ < r) \ vanishes X" unfolding vanishes_def by simp -lemma vanishesD: "\vanishes X; 0 < r\ \ \k. \n\k. \X n\ < r" +lemma vanishesD: "vanishes X \ 0 < r \ \k. \n\k. \X n\ < r" unfolding vanishes_def by simp lemma vanishes_const [simp]: "vanishes (\n. c) \ c = 0" unfolding vanishes_def - apply (cases "c = 0", auto) - apply (rule exI [where x="\c\"], auto) + apply (cases "c = 0") + apply auto + apply (rule exI [where x = "\c\"]) + apply auto done lemma vanishes_minus: "vanishes X \ vanishes (\n. - X n)" unfolding vanishes_def by simp lemma vanishes_add: - assumes X: "vanishes X" and Y: "vanishes Y" + assumes X: "vanishes X" + and Y: "vanishes Y" shows "vanishes (\n. X n + Y n)" proof (rule vanishesI) - fix r :: rat assume "0 < r" + fix r :: rat + assume "0 < r" then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" by (rule obtain_pos_sum) obtain i where i: "\n\i. \X n\ < s" @@ -93,26 +90,28 @@ obtain j where j: "\n\j. \Y n\ < t" using vanishesD [OF Y t] .. have "\n\max i j. \X n + Y n\ < r" - proof (clarsimp) - fix n assume n: "i \ n" "j \ n" + proof clarsimp + fix n + assume n: "i \ n" "j \ n" have "\X n + Y n\ \ \X n\ + \Y n\" by (rule abs_triangle_ineq) also have "\ < s + t" by (simp add: add_strict_mono i j n) finally show "\X n + Y n\ < r" unfolding r . qed - thus "\k. \n\k. \X n + Y n\ < r" .. + then show "\k. \n\k. \X n + Y n\ < r" .. qed lemma vanishes_diff: - assumes X: "vanishes X" and Y: "vanishes Y" + assumes "vanishes X" "vanishes Y" shows "vanishes (\n. X n - Y n)" - unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus X Y) + unfolding diff_conv_add_uminus by (intro vanishes_add vanishes_minus assms) lemma vanishes_mult_bounded: assumes X: "\a>0. \n. \X n\ < a" assumes Y: "vanishes (\n. Y n)" shows "vanishes (\n. X n * Y n)" proof (rule vanishesI) - fix r :: rat assume r: "0 < r" + fix r :: rat + assume r: "0 < r" obtain a where a: "0 < a" "\n. \X n\ < a" using X by blast obtain b where b: "0 < b" "r = a * b" @@ -124,22 +123,19 @@ using vanishesD [OF Y b(1)] .. have "\n\k. \X n * Y n\ < r" by (simp add: b(2) abs_mult mult_strict_mono' a k) - thus "\k. \n\k. \X n * Y n\ < r" .. + then show "\k. \n\k. \X n * Y n\ < r" .. qed + subsection \Cauchy sequences\ -definition - cauchy :: "(nat \ rat) \ bool" -where - "cauchy X \ (\r>0. \k. \m\k. \n\k. \X m - X n\ < r)" +definition cauchy :: "(nat \ rat) \ bool" + where "cauchy X \ (\r>0. \k. \m\k. \n\k. \X m - X n\ < r)" -lemma cauchyI: - "(\r. 0 < r \ \k. \m\k. \n\k. \X m - X n\ < r) \ cauchy X" +lemma cauchyI: "(\r. 0 < r \ \k. \m\k. \n\k. \X m - X n\ < r) \ cauchy X" unfolding cauchy_def by simp -lemma cauchyD: - "\cauchy X; 0 < r\ \ \k. \m\k. \n\k. \X m - X n\ < r" +lemma cauchyD: "cauchy X \ 0 < r \ \k. \m\k. \n\k. \X m - X n\ < r" unfolding cauchy_def by simp lemma cauchy_const [simp]: "cauchy (\n. x)" @@ -149,7 +145,8 @@ assumes X: "cauchy X" and Y: "cauchy Y" shows "cauchy (\n. X n + Y n)" proof (rule cauchyI) - fix r :: rat assume "0 < r" + fix r :: rat + assume "0 < r" then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" by (rule obtain_pos_sum) obtain i where i: "\m\i. \n\i. \X m - X n\ < s" @@ -157,30 +154,32 @@ obtain j where j: "\m\j. \n\j. \Y m - Y n\ < t" using cauchyD [OF Y t] .. have "\m\max i j. \n\max i j. \(X m + Y m) - (X n + Y n)\ < r" - proof (clarsimp) - fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" + proof clarsimp + fix m n + assume *: "i \ m" "j \ m" "i \ n" "j \ n" have "\(X m + Y m) - (X n + Y n)\ \ \X m - X n\ + \Y m - Y n\" unfolding add_diff_add by (rule abs_triangle_ineq) also have "\ < s + t" - by (rule add_strict_mono, simp_all add: i j *) - finally show "\(X m + Y m) - (X n + Y n)\ < r" unfolding r . + by (rule add_strict_mono) (simp_all add: i j *) + finally show "\(X m + Y m) - (X n + Y n)\ < r" by (simp only: r) qed - thus "\k. \m\k. \n\k. \(X m + Y m) - (X n + Y n)\ < r" .. + then show "\k. \m\k. \n\k. \(X m + Y m) - (X n + Y n)\ < r" .. qed lemma cauchy_minus [simp]: assumes X: "cauchy X" shows "cauchy (\n. - X n)" -using assms unfolding cauchy_def -unfolding minus_diff_minus abs_minus_cancel . + using assms unfolding cauchy_def + unfolding minus_diff_minus abs_minus_cancel . lemma cauchy_diff [simp]: - assumes X: "cauchy X" and Y: "cauchy Y" + assumes "cauchy X" "cauchy Y" shows "cauchy (\n. X n - Y n)" using assms unfolding diff_conv_add_uminus by (simp del: add_uminus_conv_diff) lemma cauchy_imp_bounded: - assumes "cauchy X" shows "\b>0. \n. \X n\ < b" + assumes "cauchy X" + shows "\b>0. \n. \X n\ < b" proof - obtain k where k: "\m\k. \n\k. \X m - X n\ < 1" using cauchyD [OF assms zero_less_one] .. @@ -189,21 +188,21 @@ have "0 \ \X 0\" by simp also have "\X 0\ \ Max (abs ` X ` {..k})" by simp finally have "0 \ Max (abs ` X ` {..k})" . - thus "0 < Max (abs ` X ` {..k}) + 1" by simp + then show "0 < Max (abs ` X ` {..k}) + 1" by simp next fix n :: nat show "\X n\ < Max (abs ` X ` {..k}) + 1" proof (rule linorder_le_cases) assume "n \ k" - hence "\X n\ \ Max (abs ` X ` {..k})" by simp - thus "\X n\ < Max (abs ` X ` {..k}) + 1" by simp + then have "\X n\ \ Max (abs ` X ` {..k})" by simp + then show "\X n\ < Max (abs ` X ` {..k}) + 1" by simp next assume "k \ n" have "\X n\ = \X k + (X n - X k)\" by simp also have "\X k + (X n - X k)\ \ \X k\ + \X n - X k\" by (rule abs_triangle_ineq) also have "\ < Max (abs ` X ` {..k}) + 1" - by (rule add_le_less_mono, simp, simp add: k \k \ n\) + by (rule add_le_less_mono) (simp_all add: k \k \ n\) finally show "\X n\ < Max (abs ` X ` {..k}) + 1" . qed qed @@ -232,8 +231,9 @@ obtain j where j: "\m\j. \n\j. \Y m - Y n\ < t" using cauchyD [OF Y t] .. have "\m\max i j. \n\max i j. \X m * Y m - X n * Y n\ < r" - proof (clarsimp) - fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" + proof clarsimp + fix m n + assume *: "i \ m" "j \ m" "i \ n" "j \ n" have "\X m * Y m - X n * Y n\ = \X m * (Y m - Y n) + (X m - X n) * Y n\" unfolding mult_diff_mult .. also have "\ \ \X m * (Y m - Y n)\ + \(X m - X n) * Y n\" @@ -242,9 +242,9 @@ unfolding abs_mult .. also have "\ < a * t + s * b" by (simp_all add: add_strict_mono mult_strict_mono' a b i j *) - finally show "\X m * Y m - X n * Y n\ < r" unfolding r . + finally show "\X m * Y m - X n * Y n\ < r" by (simp only: r) qed - thus "\k. \m\k. \n\k. \X m * Y m - X n * Y n\ < r" .. + then show "\k. \m\k. \n\k. \X m * Y m - X n * Y n\ < r" .. qed lemma cauchy_not_vanishes_cases: @@ -264,10 +264,10 @@ using i \i \ k\ by auto have "X k \ - r \ r \ X k" using \r \ \X k\\ by auto - hence "(\n\k. t < - X n) \ (\n\k. t < X n)" + then have "(\n\k. t < - X n) \ (\n\k. t < X n)" unfolding \r = s + t\ using k by auto - hence "\k. (\n\k. t < - X n) \ (\n\k. t < X n)" .. - thus "\t>0. \k. (\n\k. t < - X n) \ (\n\k. t < X n)" + then have "\k. (\n\k. t < - X n) \ (\n\k. t < X n)" .. + then show "\t>0. \k. (\n\k. t < - X n) \ (\n\k. t < X n)" using t by auto qed @@ -275,15 +275,21 @@ assumes X: "cauchy X" assumes nz: "\ vanishes X" shows "\b>0. \k. \n\k. b < \X n\" -using cauchy_not_vanishes_cases [OF assms] -by clarify (rule exI, erule conjI, rule_tac x=k in exI, auto) + using cauchy_not_vanishes_cases [OF assms] + apply clarify + apply (rule exI) + apply (erule conjI) + apply (rule_tac x = k in exI) + apply auto + done lemma cauchy_inverse [simp]: assumes X: "cauchy X" assumes nz: "\ vanishes X" shows "cauchy (\n. inverse (X n))" proof (rule cauchyI) - fix r :: rat assume "0 < r" + fix r :: rat + assume "0 < r" obtain b i where b: "0 < b" and i: "\n\i. b < \X n\" using cauchy_not_vanishes [OF X nz] by blast from b i have nz: "\n\i. X n \ 0" by auto @@ -296,84 +302,84 @@ obtain j where j: "\m\j. \n\j. \X m - X n\ < s" using cauchyD [OF X s] .. have "\m\max i j. \n\max i j. \inverse (X m) - inverse (X n)\ < r" - proof (clarsimp) - fix m n assume *: "i \ m" "j \ m" "i \ n" "j \ n" - have "\inverse (X m) - inverse (X n)\ = - inverse \X m\ * \X m - X n\ * inverse \X n\" + proof clarsimp + fix m n + assume *: "i \ m" "j \ m" "i \ n" "j \ n" + have "\inverse (X m) - inverse (X n)\ = inverse \X m\ * \X m - X n\ * inverse \X n\" by (simp add: inverse_diff_inverse nz * abs_mult) also have "\ < inverse b * s * inverse b" - by (simp add: mult_strict_mono less_imp_inverse_less - i j b * s) - finally show "\inverse (X m) - inverse (X n)\ < r" unfolding r . + by (simp add: mult_strict_mono less_imp_inverse_less i j b * s) + finally show "\inverse (X m) - inverse (X n)\ < r" by (simp only: r) qed - thus "\k. \m\k. \n\k. \inverse (X m) - inverse (X n)\ < r" .. + then show "\k. \m\k. \n\k. \inverse (X m) - inverse (X n)\ < r" .. qed lemma vanishes_diff_inverse: assumes X: "cauchy X" "\ vanishes X" - assumes Y: "cauchy Y" "\ vanishes Y" - assumes XY: "vanishes (\n. X n - Y n)" + and Y: "cauchy Y" "\ vanishes Y" + and XY: "vanishes (\n. X n - Y n)" shows "vanishes (\n. inverse (X n) - inverse (Y n))" proof (rule vanishesI) - fix r :: rat assume r: "0 < r" + fix r :: rat + assume r: "0 < r" obtain a i where a: "0 < a" and i: "\n\i. a < \X n\" using cauchy_not_vanishes [OF X] by blast obtain b j where b: "0 < b" and j: "\n\j. b < \Y n\" using cauchy_not_vanishes [OF Y] by blast obtain s where s: "0 < s" and "inverse a * s * inverse b = r" proof - show "0 < a * r * b" - using a r b by simp - show "inverse a * (a * r * b) * inverse b = r" - using a r b by simp + show "0 < a * r * b" using a r b by simp + show "inverse a * (a * r * b) * inverse b = r" using a r b by simp qed obtain k where k: "\n\k. \X n - Y n\ < s" using vanishesD [OF XY s] .. have "\n\max (max i j) k. \inverse (X n) - inverse (Y n)\ < r" - proof (clarsimp) - fix n assume n: "i \ n" "j \ n" "k \ n" - have "X n \ 0" and "Y n \ 0" - using i j a b n by auto - hence "\inverse (X n) - inverse (Y n)\ = - inverse \X n\ * \X n - Y n\ * inverse \Y n\" + proof clarsimp + fix n + assume n: "i \ n" "j \ n" "k \ n" + with i j a b have "X n \ 0" and "Y n \ 0" + by auto + then have "\inverse (X n) - inverse (Y n)\ = inverse \X n\ * \X n - Y n\ * inverse \Y n\" by (simp add: inverse_diff_inverse abs_mult) also have "\ < inverse a * s * inverse b" - apply (intro mult_strict_mono' less_imp_inverse_less) - apply (simp_all add: a b i j k n) - done + by (intro mult_strict_mono' less_imp_inverse_less) (simp_all add: a b i j k n) also note \inverse a * s * inverse b = r\ finally show "\inverse (X n) - inverse (Y n)\ < r" . qed - thus "\k. \n\k. \inverse (X n) - inverse (Y n)\ < r" .. + then show "\k. \n\k. \inverse (X n) - inverse (Y n)\ < r" .. qed + subsection \Equivalence relation on Cauchy sequences\ definition realrel :: "(nat \ rat) \ (nat \ rat) \ bool" where "realrel = (\X Y. cauchy X \ cauchy Y \ vanishes (\n. X n - Y n))" -lemma realrelI [intro?]: - assumes "cauchy X" and "cauchy Y" and "vanishes (\n. X n - Y n)" - shows "realrel X Y" - using assms unfolding realrel_def by simp +lemma realrelI [intro?]: "cauchy X \ cauchy Y \ vanishes (\n. X n - Y n) \ realrel X Y" + by (simp add: realrel_def) lemma realrel_refl: "cauchy X \ realrel X X" - unfolding realrel_def by simp + by (simp add: realrel_def) lemma symp_realrel: "symp realrel" unfolding realrel_def - by (rule sympI, clarify, drule vanishes_minus, simp) + apply (rule sympI) + apply clarify + apply (drule vanishes_minus) + apply simp + done lemma transp_realrel: "transp realrel" unfolding realrel_def - apply (rule transpI, clarify) + apply (rule transpI) + apply clarify apply (drule (1) vanishes_add) apply (simp add: algebra_simps) done lemma part_equivp_realrel: "part_equivp realrel" - by (blast intro: part_equivpI symp_realrel transp_realrel - realrel_refl cauchy_const) + by (blast intro: part_equivpI symp_realrel transp_realrel realrel_refl cauchy_const) + subsection \The field of real numbers\ @@ -385,20 +391,20 @@ unfolding real.pcr_cr_eq cr_real_def realrel_def by auto lemma Real_induct [induct type: real]: (* TODO: generate automatically *) - assumes "\X. cauchy X \ P (Real X)" shows "P x" + assumes "\X. cauchy X \ P (Real X)" + shows "P x" proof (induct x) case (1 X) - hence "cauchy X" by (simp add: realrel_def) - thus "P (Real X)" by (rule assms) + then have "cauchy X" by (simp add: realrel_def) + then show "P (Real X)" by (rule assms) qed -lemma eq_Real: - "cauchy X \ cauchy Y \ Real X = Real Y \ vanishes (\n. X n - Y n)" +lemma eq_Real: "cauchy X \ cauchy Y \ Real X = Real Y \ vanishes (\n. X n - Y n)" using real.rel_eq_transfer unfolding real.pcr_cr_eq cr_real_def rel_fun_def realrel_def by simp lemma Domainp_pcr_real [transfer_domain_rule]: "Domainp pcr_real = cauchy" -by (simp add: real.domain_eq realrel_def) + by (simp add: real.domain_eq realrel_def) instantiation real :: field begin @@ -419,14 +425,16 @@ lift_definition times_real :: "real \ real \ real" is "\X Y n. X n * Y n" unfolding realrel_def mult_diff_mult - by (subst (4) mult.commute, simp only: cauchy_mult vanishes_add - vanishes_mult_bounded cauchy_imp_bounded simp_thms) + apply (subst (4) mult.commute) + apply (simp only: cauchy_mult vanishes_add vanishes_mult_bounded cauchy_imp_bounded simp_thms) + done lift_definition inverse_real :: "real \ real" is "\X. if vanishes X then (\n. 0) else (\n. inverse (X n))" proof - - fix X Y assume "realrel X Y" - hence X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\n. X n - Y n)" + fix X Y + assume "realrel X Y" + then have X: "cauchy X" and Y: "cauchy Y" and XY: "vanishes (\n. X n - Y n)" unfolding realrel_def by simp_all have "vanishes X \ vanishes Y" proof @@ -436,49 +444,32 @@ assume "vanishes Y" from vanishes_add [OF this XY] show "vanishes X" by simp qed - thus "?thesis X Y" - unfolding realrel_def - by (simp add: vanishes_diff_inverse X Y XY) + then show "?thesis X Y" by (simp add: vanishes_diff_inverse X Y XY realrel_def) qed -definition - "x - y = (x::real) + - y" - -definition - "x div y = (x::real) * inverse y" +definition "x - y = x + - y" for x y :: real -lemma add_Real: - assumes X: "cauchy X" and Y: "cauchy Y" - shows "Real X + Real Y = Real (\n. X n + Y n)" - using assms plus_real.transfer - unfolding cr_real_eq rel_fun_def by simp +definition "x div y = x * inverse y" for x y :: real + +lemma add_Real: "cauchy X \ cauchy Y \ Real X + Real Y = Real (\n. X n + Y n)" + using plus_real.transfer by (simp add: cr_real_eq rel_fun_def) -lemma minus_Real: - assumes X: "cauchy X" - shows "- Real X = Real (\n. - X n)" - using assms uminus_real.transfer - unfolding cr_real_eq rel_fun_def by simp +lemma minus_Real: "cauchy X \ - Real X = Real (\n. - X n)" + using uminus_real.transfer by (simp add: cr_real_eq rel_fun_def) -lemma diff_Real: - assumes X: "cauchy X" and Y: "cauchy Y" - shows "Real X - Real Y = Real (\n. X n - Y n)" - unfolding minus_real_def - by (simp add: minus_Real add_Real X Y) +lemma diff_Real: "cauchy X \ cauchy Y \ Real X - Real Y = Real (\n. X n - Y n)" + by (simp add: minus_Real add_Real minus_real_def) -lemma mult_Real: - assumes X: "cauchy X" and Y: "cauchy Y" - shows "Real X * Real Y = Real (\n. X n * Y n)" - using assms times_real.transfer - unfolding cr_real_eq rel_fun_def by simp +lemma mult_Real: "cauchy X \ cauchy Y \ Real X * Real Y = Real (\n. X n * Y n)" + using times_real.transfer by (simp add: cr_real_eq rel_fun_def) lemma inverse_Real: - assumes X: "cauchy X" - shows "inverse (Real X) = - (if vanishes X then 0 else Real (\n. inverse (X n)))" - using assms inverse_real.transfer zero_real.transfer + "cauchy X \ inverse (Real X) = (if vanishes X then 0 else Real (\n. inverse (X n)))" + using inverse_real.transfer zero_real.transfer unfolding cr_real_eq rel_fun_def by (simp split: if_split_asm, metis) -instance proof +instance +proof fix a b c :: real show "a + b = b + a" by transfer (simp add: ac_simps realrel_def) @@ -516,115 +507,125 @@ end + subsection \Positive reals\ lift_definition positive :: "real \ bool" is "\X. \r>0. \k. \n\k. r < X n" proof - - { fix X Y - assume "realrel X Y" - hence XY: "vanishes (\n. X n - Y n)" - unfolding realrel_def by simp_all - assume "\r>0. \k. \n\k. r < X n" - then obtain r i where "0 < r" and i: "\n\i. r < X n" + have 1: "\r>0. \k. \n\k. r < Y n" + if *: "realrel X Y" and **: "\r>0. \k. \n\k. r < X n" for X Y + proof - + from * have XY: "vanishes (\n. X n - Y n)" + by (simp_all add: realrel_def) + from ** obtain r i where "0 < r" and i: "\n\i. r < X n" by blast obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" using \0 < r\ by (rule obtain_pos_sum) obtain j where j: "\n\j. \X n - Y n\ < s" using vanishesD [OF XY s] .. have "\n\max i j. t < Y n" - proof (clarsimp) - fix n assume n: "i \ n" "j \ n" + proof clarsimp + fix n + assume n: "i \ n" "j \ n" have "\X n - Y n\ < s" and "r < X n" using i j n by simp_all - thus "t < Y n" unfolding r by simp + then show "t < Y n" by (simp add: r) qed - hence "\r>0. \k. \n\k. r < Y n" using t by blast - } note 1 = this + then show ?thesis using t by blast + qed fix X Y assume "realrel X Y" - hence "realrel X Y" and "realrel Y X" - using symp_realrel unfolding symp_def by auto - thus "?thesis X Y" + then have "realrel X Y" and "realrel Y X" + using symp_realrel by (auto simp: symp_def) + then show "?thesis X Y" by (safe elim!: 1) qed -lemma positive_Real: - assumes X: "cauchy X" - shows "positive (Real X) \ (\r>0. \k. \n\k. r < X n)" - using assms positive.transfer - unfolding cr_real_eq rel_fun_def by simp +lemma positive_Real: "cauchy X \ positive (Real X) \ (\r>0. \k. \n\k. r < X n)" + using positive.transfer by (simp add: cr_real_eq rel_fun_def) lemma positive_zero: "\ positive 0" by transfer auto -lemma positive_add: - "positive x \ positive y \ positive (x + y)" -apply transfer -apply (clarify, rename_tac a b i j) -apply (rule_tac x="a + b" in exI, simp) -apply (rule_tac x="max i j" in exI, clarsimp) -apply (simp add: add_strict_mono) -done +lemma positive_add: "positive x \ positive y \ positive (x + y)" + apply transfer + apply clarify + apply (rename_tac a b i j) + apply (rule_tac x = "a + b" in exI) + apply simp + apply (rule_tac x = "max i j" in exI) + apply clarsimp + apply (simp add: add_strict_mono) + done -lemma positive_mult: - "positive x \ positive y \ positive (x * y)" -apply transfer -apply (clarify, rename_tac a b i j) -apply (rule_tac x="a * b" in exI, simp) -apply (rule_tac x="max i j" in exI, clarsimp) -apply (rule mult_strict_mono, auto) -done +lemma positive_mult: "positive x \ positive y \ positive (x * y)" + apply transfer + apply clarify + apply (rename_tac a b i j) + apply (rule_tac x = "a * b" in exI) + apply simp + apply (rule_tac x = "max i j" in exI) + apply clarsimp + apply (rule mult_strict_mono) + apply auto + done -lemma positive_minus: - "\ positive x \ x \ 0 \ positive (- x)" -apply transfer -apply (simp add: realrel_def) -apply (drule (1) cauchy_not_vanishes_cases, safe) -apply blast+ -done +lemma positive_minus: "\ positive x \ x \ 0 \ positive (- x)" + apply transfer + apply (simp add: realrel_def) + apply (drule (1) cauchy_not_vanishes_cases, safe) + apply blast+ + done instantiation real :: linordered_field begin -definition - "x < y \ positive (y - x)" +definition "x < y \ positive (y - x)" -definition - "x \ (y::real) \ x < y \ x = y" +definition "x \ y \ x < y \ x = y" for x y :: real -definition - "\a::real\ = (if a < 0 then - a else a)" +definition "\a\ = (if a < 0 then - a else a)" for a :: real -definition - "sgn (a::real) = (if a = 0 then 0 else if 0 < a then 1 else - 1)" +definition "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" for a :: real -instance proof +instance +proof fix a b c :: real show "\a\ = (if a < 0 then - a else a)" by (rule abs_real_def) show "a < b \ a \ b \ \ b \ a" unfolding less_eq_real_def less_real_def - by (auto, drule (1) positive_add, simp_all add: positive_zero) + apply auto + apply (drule (1) positive_add) + apply (simp_all add: positive_zero) + done show "a \ a" unfolding less_eq_real_def by simp show "a \ b \ b \ c \ a \ c" unfolding less_eq_real_def less_real_def - by (auto, drule (1) positive_add, simp add: algebra_simps) + apply auto + apply (drule (1) positive_add) + apply (simp add: algebra_simps) + done show "a \ b \ b \ a \ a = b" unfolding less_eq_real_def less_real_def - by (auto, drule (1) positive_add, simp add: positive_zero) + apply auto + apply (drule (1) positive_add) + apply (simp add: positive_zero) + done show "a \ b \ c + a \ c + b" - unfolding less_eq_real_def less_real_def by auto + by (auto simp: less_eq_real_def less_real_def) (* FIXME: Procedure int_combine_numerals: c + b - (c + a) \ b + - a *) (* Should produce c + b - (c + a) \ b - a *) show "sgn a = (if a = 0 then 0 else if 0 < a then 1 else - 1)" by (rule sgn_real_def) show "a \ b \ b \ a" - unfolding less_eq_real_def less_real_def - by (auto dest!: positive_minus) + by (auto dest!: positive_minus simp: less_eq_real_def less_real_def) show "a < b \ 0 < c \ c * a < c * b" unfolding less_real_def - by (drule (1) positive_mult, simp add: algebra_simps) + apply (drule (1) positive_mult) + apply (simp add: algebra_simps) + done qed end @@ -632,34 +633,26 @@ instantiation real :: distrib_lattice begin -definition - "(inf :: real \ real \ real) = min" +definition "(inf :: real \ real \ real) = min" -definition - "(sup :: real \ real \ real) = max" +definition "(sup :: real \ real \ real) = max" -instance proof -qed (auto simp add: inf_real_def sup_real_def max_min_distrib2) +instance by standard (auto simp add: inf_real_def sup_real_def max_min_distrib2) end lemma of_nat_Real: "of_nat x = Real (\n. of_nat x)" -apply (induct x) -apply (simp add: zero_real_def) -apply (simp add: one_real_def add_Real) -done + by (induct x) (simp_all add: zero_real_def one_real_def add_Real) lemma of_int_Real: "of_int x = Real (\n. of_int x)" -apply (cases x rule: int_diff_cases) -apply (simp add: of_nat_Real diff_Real) -done + by (cases x rule: int_diff_cases) (simp add: of_nat_Real diff_Real) lemma of_rat_Real: "of_rat x = Real (\n. x)" -apply (induct x) -apply (simp add: Fract_of_int_quotient of_rat_divide) -apply (simp add: of_int_Real divide_inverse) -apply (simp add: inverse_Real mult_Real) -done + apply (induct x) + apply (simp add: Fract_of_int_quotient of_rat_divide) + apply (simp add: of_int_Real divide_inverse) + apply (simp add: inverse_Real mult_Real) + done instance real :: archimedean_field proof @@ -681,69 +674,82 @@ instantiation real :: floor_ceiling begin -definition [code del]: - "\x::real\ = (THE z. of_int z \ x \ x < of_int (z + 1))" +definition [code del]: "\x::real\ = (THE z. of_int z \ x \ x < of_int (z + 1))" instance proof - fix x :: real - show "of_int \x\ \ x \ x < of_int (\x\ + 1)" + show "of_int \x\ \ x \ x < of_int (\x\ + 1)" for x :: real unfolding floor_real_def using floor_exists1 by (rule theI') qed end + subsection \Completeness\ lemma not_positive_Real: assumes X: "cauchy X" shows "\ positive (Real X) \ (\r>0. \k. \n\k. X n \ r)" -unfolding positive_Real [OF X] -apply (auto, unfold not_less) -apply (erule obtain_pos_sum) -apply (drule_tac x=s in spec, simp) -apply (drule_tac r=t in cauchyD [OF X], clarify) -apply (drule_tac x=k in spec, clarsimp) -apply (rule_tac x=n in exI, clarify, rename_tac m) -apply (drule_tac x=m in spec, simp) -apply (drule_tac x=n in spec, simp) -apply (drule spec, drule (1) mp, clarify, rename_tac i) -apply (rule_tac x="max i k" in exI, simp) -done + unfolding positive_Real [OF X] + apply auto + apply (unfold not_less) + apply (erule obtain_pos_sum) + apply (drule_tac x=s in spec) + apply simp + apply (drule_tac r=t in cauchyD [OF X]) + apply clarify + apply (drule_tac x=k in spec) + apply clarsimp + apply (rule_tac x=n in exI) + apply clarify + apply (rename_tac m) + apply (drule_tac x=m in spec) + apply simp + apply (drule_tac x=n in spec) + apply simp + apply (drule spec) + apply (drule (1) mp) + apply clarify + apply (rename_tac i) + apply (rule_tac x = "max i k" in exI) + apply simp + done lemma le_Real: - assumes X: "cauchy X" and Y: "cauchy Y" + assumes "cauchy X" "cauchy Y" shows "Real X \ Real Y = (\r>0. \k. \n\k. X n \ Y n + r)" -unfolding not_less [symmetric, where 'a=real] less_real_def -apply (simp add: diff_Real not_positive_Real X Y) -apply (simp add: diff_le_eq ac_simps) -done + unfolding not_less [symmetric, where 'a=real] less_real_def + apply (simp add: diff_Real not_positive_Real assms) + apply (simp add: diff_le_eq ac_simps) + done lemma le_RealI: assumes Y: "cauchy Y" shows "\n. x \ of_rat (Y n) \ x \ Real Y" proof (induct x) - fix X assume X: "cauchy X" and "\n. Real X \ of_rat (Y n)" - hence le: "\m r. 0 < r \ \k. \n\k. X n \ Y m + r" + fix X + assume X: "cauchy X" and "\n. Real X \ of_rat (Y n)" + then have le: "\m r. 0 < r \ \k. \n\k. X n \ Y m + r" by (simp add: of_rat_Real le_Real) - { - fix r :: rat assume "0 < r" - then obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" + then have "\k. \n\k. X n \ Y n + r" if "0 < r" for r :: rat + proof - + from that obtain s t where s: "0 < s" and t: "0 < t" and r: "r = s + t" by (rule obtain_pos_sum) obtain i where i: "\m\i. \n\i. \Y m - Y n\ < s" using cauchyD [OF Y s] .. obtain j where j: "\n\j. X n \ Y i + t" using le [OF t] .. have "\n\max i j. X n \ Y n + r" - proof (clarsimp) - fix n assume n: "i \ n" "j \ n" + proof clarsimp + fix n + assume n: "i \ n" "j \ n" have "X n \ Y i + t" using n j by simp moreover have "\Y i - Y n\ < s" using n i by simp ultimately show "X n \ Y n + r" unfolding r by simp qed - hence "\k. \n\k. X n \ Y n + r" .. - } - thus "Real X \ Real Y" + then show ?thesis .. + qed + then show "Real X \ Real Y" by (simp add: of_rat_Real le_Real X Y) qed @@ -754,18 +760,22 @@ proof - have "- y \ - Real X" by (simp add: minus_Real X le_RealI of_rat_minus le) - thus ?thesis by simp + then show ?thesis by simp qed lemma less_RealD: - assumes Y: "cauchy Y" + assumes "cauchy Y" shows "x < Real Y \ \n. x < of_rat (Y n)" -by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y]) + apply (erule contrapos_pp) + apply (simp add: not_less) + apply (erule Real_leI [OF assms]) + done -lemma of_nat_less_two_power [simp]: - "of_nat n < (2::'a::linordered_idom) ^ n" -apply (induct n, simp) -by (metis add_le_less_mono mult_2 of_nat_Suc one_le_numeral one_le_power power_Suc) +lemma of_nat_less_two_power [simp]: "of_nat n < (2::'a::linordered_idom) ^ n" + apply (induct n) + apply simp + apply (metis add_le_less_mono mult_2 of_nat_Suc one_le_numeral one_le_power power_Suc) + done lemma complete_real: fixes S :: "real set" @@ -781,7 +791,7 @@ have "of_int \x - 1\ \ x - 1" by (rule of_int_floor_le) also have "x - 1 < x" by simp finally have "of_int \x - 1\ < x" . - hence "\ x \ of_int \x - 1\" by (simp only: not_le) + then have "\ x \ of_int \x - 1\" by (simp only: not_le) then show "\ P (of_int \x - 1\)" unfolding P_def of_rat_of_int_eq using x by blast qed @@ -791,7 +801,7 @@ unfolding P_def of_rat_of_int_eq proof fix y assume "y \ S" - hence "y \ z" using z by simp + then have "y \ z" using z by simp also have "z \ of_int \z\" by (rule le_of_int_ceiling) finally show "y \ of_int \z\" . qed @@ -809,13 +819,13 @@ have B_Suc [simp]: "\n. B (Suc n) = (if P (C n) then C n else B n)" unfolding A_def B_def C_def bisect_def split_def by simp - have width: "\n. B n - A n = (b - a) / 2^n" - apply (simp add: eq_divide_eq) - apply (induct_tac n, simp) - apply (simp add: C_def avg_def algebra_simps) + have width: "B n - A n = (b - a) / 2^n" for n + apply (induct n) + apply (simp_all add: eq_divide_eq) + apply (simp_all add: C_def avg_def algebra_simps) done - have twos: "\y r :: rat. 0 < r \ \n. y / 2 ^ n < r" + have twos: "0 < r \ \n. y / 2 ^ n < r" for y r :: rat apply (simp add: divide_less_eq) apply (subst mult.commute) apply (frule_tac y=y in ex_less_of_nat_mult) @@ -828,10 +838,8 @@ apply assumption done - have PA: "\n. \ P (A n)" - by (induct_tac n, simp_all add: a) - have PB: "\n. P (B n)" - by (induct_tac n, simp_all add: b) + have PA: "\ P (A n)" for n by (induct n) (simp_all add: a) + have PB: "P (B n)" for n by (induct n) (simp_all add: b) have ab: "a < b" using a b unfolding P_def apply (clarsimp simp add: not_le) @@ -839,8 +847,7 @@ apply (drule (1) less_le_trans) apply (simp add: of_rat_less) done - have AB: "\n. A n < B n" - by (induct_tac n, simp add: ab, simp add: C_def avg_def) + have AB: "A n < B n" for n by (induct n) (simp_all add: ab C_def avg_def) have A_mono: "\i j. i \ j \ A i \ A j" apply (auto simp add: le_less [where 'a=nat]) apply (erule less_Suc_induct) @@ -857,8 +864,7 @@ apply (rule AB [THEN less_imp_le]) apply simp done - have cauchy_lemma: - "\X. \n. \i\n. A n \ X i \ X i \ B n \ cauchy X" + have cauchy_lemma: "\X. \n. \i\n. A n \ X i \ X i \ B n \ cauchy X" apply (rule cauchyI) apply (drule twos [where y="b - a"]) apply (erule exE) @@ -884,7 +890,8 @@ done have 1: "\x\S. x \ Real B" proof - fix x assume "x \ S" + fix x + assume "x \ S" then show "x \ Real B" using PB [unfolded P_def] \cauchy B\ by (simp add: le_RealI) @@ -902,12 +909,14 @@ done have "vanishes (\n. (b - a) / 2 ^ n)" proof (rule vanishesI) - fix r :: rat assume "0 < r" + fix r :: rat + assume "0 < r" then obtain k where k: "\b - a\ / 2 ^ k < r" using twos by blast have "\n\k. \(b - a) / 2 ^ n\ < r" - proof (clarify) - fix n assume n: "k \ n" + proof clarify + fix n + assume n: "k \ n" have "\(b - a) / 2 ^ n\ = \b - a\ / 2 ^ n" by simp also have "\ \ \b - a\ / 2 ^ k" @@ -915,52 +924,55 @@ also note k finally show "\(b - a) / 2 ^ n\ < r" . qed - thus "\k. \n\k. \(b - a) / 2 ^ n\ < r" .. + then show "\k. \n\k. \(b - a) / 2 ^ n\ < r" .. qed - hence 3: "Real B = Real A" + then have 3: "Real B = Real A" by (simp add: eq_Real \cauchy A\ \cauchy B\ width) show "\y. (\x\S. x \ y) \ (\z. (\x\S. x \ z) \ y \ z)" - using 1 2 3 by (rule_tac x="Real B" in exI, simp) + apply (rule exI [where x = "Real B"]) + using 1 2 3 + apply simp + done qed instantiation real :: linear_continuum begin -subsection\Supremum of a set of reals\ +subsection \Supremum of a set of reals\ definition "Sup X = (LEAST z::real. \x\X. x \ z)" -definition "Inf (X::real set) = - Sup (uminus ` X)" +definition "Inf X = - Sup (uminus ` X)" for X :: "real set" instance proof - { fix x :: real and X :: "real set" - assume x: "x \ X" "bdd_above X" - then obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" + show Sup_upper: "x \ Sup X" if "x \ X" "bdd_above X" for x :: real and X :: "real set" + proof - + from that obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" using complete_real[of X] unfolding bdd_above_def by blast - then show "x \ Sup X" - unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) } - note Sup_upper = this - - { fix z :: real and X :: "real set" - assume x: "X \ {}" and z: "\x. x \ X \ x \ z" - then obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" - using complete_real[of X] by blast + then show ?thesis unfolding Sup_real_def by (rule LeastI2_order) (auto simp: that) + qed + show Sup_least: "Sup X \ z" if "X \ {}" and z: "\x. x \ X \ x \ z" + for z :: real and X :: "real set" + proof - + from that obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" + using complete_real [of X] by blast then have "Sup X = s" unfolding Sup_real_def by (best intro: Least_equality) - also from s z have "... \ z" + also from s z have "\ \ z" by blast - finally show "Sup X \ z" . } - note Sup_least = this - - { fix x :: real and X :: "real set" assume x: "x \ X" "bdd_below X" then show "Inf X \ x" - using Sup_upper[of "-x" "uminus ` X"] by (auto simp: Inf_real_def) } - { fix z :: real and X :: "real set" assume "X \ {}" "\x. x \ X \ z \ x" then show "z \ Inf X" - using Sup_least[of "uminus ` X" "- z"] by (force simp: Inf_real_def) } + finally show ?thesis . + qed + show "Inf X \ x" if "x \ X" "bdd_below X" for x :: real and X :: "real set" + using Sup_upper [of "-x" "uminus ` X"] by (auto simp: Inf_real_def that) + show "z \ Inf X" if "X \ {}" "\x. x \ X \ z \ x" for z :: real and X :: "real set" + using Sup_least [of "uminus ` X" "- z"] by (force simp: Inf_real_def that) show "\a b::real. a \ b" using zero_neq_one by blast qed + end + subsection \Hiding implementation details\ hide_const (open) vanishes cauchy positive Real @@ -972,42 +984,35 @@ lifting_update real.lifting lifting_forget real.lifting -subsection\More Lemmas\ + +subsection \More Lemmas\ text \BH: These lemmas should not be necessary; they should be -covered by existing simp rules and simplification procedures.\ + covered by existing simp rules and simplification procedures.\ -lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)" -by simp (* solved by linordered_ring_less_cancel_factor simproc *) +lemma real_mult_less_iff1 [simp]: "0 < z \ x * z < y * z \ x < y" for x y z :: real + by simp (* solved by linordered_ring_less_cancel_factor simproc *) -lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \ y*z) = (x\y)" -by simp (* solved by linordered_ring_le_cancel_factor simproc *) +lemma real_mult_le_cancel_iff1 [simp]: "0 < z \ x * z \ y * z \ x \ y" for x y z :: real + by simp (* solved by linordered_ring_le_cancel_factor simproc *) -lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \ z*y) = (x\y)" -by simp (* solved by linordered_ring_le_cancel_factor simproc *) +lemma real_mult_le_cancel_iff2 [simp]: "0 < z \ z * x \ z * y \ x \ y" for x y z :: real + by simp (* solved by linordered_ring_le_cancel_factor simproc *) subsection \Embedding numbers into the Reals\ -abbreviation - real_of_nat :: "nat \ real" -where - "real_of_nat \ of_nat" +abbreviation real_of_nat :: "nat \ real" + where "real_of_nat \ of_nat" -abbreviation - real :: "nat \ real" -where - "real \ of_nat" +abbreviation real :: "nat \ real" + where "real \ of_nat" -abbreviation - real_of_int :: "int \ real" -where - "real_of_int \ of_int" +abbreviation real_of_int :: "int \ real" + where "real_of_int \ of_int" -abbreviation - real_of_rat :: "rat \ real" -where - "real_of_rat \ of_rat" +abbreviation real_of_rat :: "rat \ real" + where "real_of_rat \ of_rat" declare [[coercion_enabled]] @@ -1036,68 +1041,65 @@ declare of_int_1_less_iff [algebra, presburger] declare of_int_1_le_iff [algebra, presburger] -lemma int_less_real_le: "(n < m) = (real_of_int n + 1 <= real_of_int m)" +lemma int_less_real_le: "n < m \ real_of_int n + 1 \ real_of_int m" proof - have "(0::real) \ 1" by (metis less_eq_real_def zero_less_one) - thus ?thesis + then show ?thesis by (metis floor_of_int less_floor_iff) qed -lemma int_le_real_less: "(n \ m) = (real_of_int n < real_of_int m + 1)" +lemma int_le_real_less: "n \ m \ real_of_int n < real_of_int m + 1" by (meson int_less_real_le not_le) - -lemma real_of_int_div_aux: "(real_of_int x) / (real_of_int d) = +lemma real_of_int_div_aux: + "(real_of_int x) / (real_of_int d) = real_of_int (x div d) + (real_of_int (x mod d)) / (real_of_int d)" proof - have "x = (x div d) * d + x mod d" by auto then have "real_of_int x = real_of_int (x div d) * real_of_int d + real_of_int(x mod d)" by (metis of_int_add of_int_mult) - then have "real_of_int x / real_of_int d = ... / real_of_int d" + then have "real_of_int x / real_of_int d = \ / real_of_int d" by simp then show ?thesis by (auto simp add: add_divide_distrib algebra_simps) qed lemma real_of_int_div: - fixes d n :: int - shows "d dvd n \ real_of_int (n div d) = real_of_int n / real_of_int d" + "d dvd n \ real_of_int (n div d) = real_of_int n / real_of_int d" for d n :: int by (simp add: real_of_int_div_aux) -lemma real_of_int_div2: - "0 <= real_of_int n / real_of_int x - real_of_int (n div x)" - apply (case_tac "x = 0", simp) - apply (case_tac "0 < x") +lemma real_of_int_div2: "0 \ real_of_int n / real_of_int x - real_of_int (n div x)" + apply (cases "x = 0") + apply simp + apply (cases "0 < x") apply (metis add.left_neutral floor_correct floor_divide_of_int_eq le_diff_eq) apply (metis add.left_neutral floor_correct floor_divide_of_int_eq le_diff_eq) done -lemma real_of_int_div3: - "real_of_int (n::int) / real_of_int (x) - real_of_int (n div x) <= 1" +lemma real_of_int_div3: "real_of_int n / real_of_int x - real_of_int (n div x) \ 1" apply (simp add: algebra_simps) apply (subst real_of_int_div_aux) apply (auto simp add: divide_le_eq intro: order_less_imp_le) -done + done -lemma real_of_int_div4: "real_of_int (n div x) <= real_of_int (n::int) / real_of_int x" -by (insert real_of_int_div2 [of n x], simp) +lemma real_of_int_div4: "real_of_int (n div x) \ real_of_int n / real_of_int x" + using real_of_int_div2 [of n x] by simp -subsection\Embedding the Naturals into the Reals\ +subsection \Embedding the Naturals into the Reals\ -lemma real_of_card: "real (card A) = setsum (\x.1) A" +lemma real_of_card: "real (card A) = setsum (\x. 1) A" by simp -lemma nat_less_real_le: "(n < m) = (real n + 1 \ real m)" +lemma nat_less_real_le: "n < m \ real n + 1 \ real m" by (metis discrete of_nat_1 of_nat_add of_nat_le_iff) -lemma nat_le_real_less: "((n::nat) <= m) = (real n < real m + 1)" +lemma nat_le_real_less: "n \ m \ real n < real m + 1" for m n :: nat by (meson nat_less_real_le not_le) -lemma real_of_nat_div_aux: "(real x) / (real d) = - real (x div d) + (real (x mod d)) / (real d)" +lemma real_of_nat_div_aux: "real x / real d = real (x div d) + real (x mod d) / real d" proof - have "x = (x div d) * d + x mod d" by auto @@ -1110,27 +1112,25 @@ qed lemma real_of_nat_div: "d dvd n \ real(n div d) = real n / real d" - by (subst real_of_nat_div_aux) - (auto simp add: dvd_eq_mod_eq_0 [symmetric]) + by (subst real_of_nat_div_aux) (auto simp add: dvd_eq_mod_eq_0 [symmetric]) -lemma real_of_nat_div2: - "0 <= real (n::nat) / real (x) - real (n div x)" -apply (simp add: algebra_simps) -apply (subst real_of_nat_div_aux) -apply simp -done +lemma real_of_nat_div2: "0 \ real n / real x - real (n div x)" for n x :: nat + apply (simp add: algebra_simps) + apply (subst real_of_nat_div_aux) + apply simp + done -lemma real_of_nat_div3: - "real (n::nat) / real (x) - real (n div x) <= 1" -apply(case_tac "x = 0") -apply (simp) -apply (simp add: algebra_simps) -apply (subst real_of_nat_div_aux) -apply simp -done +lemma real_of_nat_div3: "real n / real x - real (n div x) \ 1" for n x :: nat + apply (cases "x = 0") + apply simp + apply (simp add: algebra_simps) + apply (subst real_of_nat_div_aux) + apply simp + done -lemma real_of_nat_div4: "real (n div x) <= real (n::nat) / real x" -by (insert real_of_nat_div2 [of n x], simp) +lemma real_of_nat_div4: "real (n div x) \ real n / real x" for n x :: nat + using real_of_nat_div2 [of n x] by simp + subsection \The Archimedean Property of the Reals\ @@ -1145,77 +1145,85 @@ lemma real_archimedian_rdiv_eq_0: assumes x0: "x \ 0" - and c: "c \ 0" - and xc: "\m::nat. m > 0 \ real m * x \ c" - shows "x = 0" -by (metis reals_Archimedean3 dual_order.order_iff_strict le0 le_less_trans not_le x0 xc) + and c: "c \ 0" + and xc: "\m::nat. m > 0 \ real m * x \ c" + shows "x = 0" + by (metis reals_Archimedean3 dual_order.order_iff_strict le0 le_less_trans not_le x0 xc) -subsection\Rationals\ +subsection \Rationals\ -lemma Rats_eq_int_div_int: - "\ = { real_of_int i / real_of_int j |i j. j \ 0}" (is "_ = ?S") +lemma Rats_eq_int_div_int: "\ = {real_of_int i / real_of_int j | i j. j \ 0}" (is "_ = ?S") proof show "\ \ ?S" proof - fix x::real assume "x : \" - then obtain r where "x = of_rat r" unfolding Rats_def .. - have "of_rat r : ?S" - by (cases r) (auto simp add:of_rat_rat) - thus "x : ?S" using \x = of_rat r\ by simp + fix x :: real + assume "x \ \" + then obtain r where "x = of_rat r" + unfolding Rats_def .. + have "of_rat r \ ?S" + by (cases r) (auto simp add: of_rat_rat) + then show "x \ ?S" + using \x = of_rat r\ by simp qed next show "?S \ \" - proof(auto simp:Rats_def) - fix i j :: int assume "j \ 0" - hence "real_of_int i / real_of_int j = of_rat(Fract i j)" + proof (auto simp: Rats_def) + fix i j :: int + assume "j \ 0" + then have "real_of_int i / real_of_int j = of_rat (Fract i j)" by (simp add: of_rat_rat) - thus "real_of_int i / real_of_int j \ range of_rat" by blast + then show "real_of_int i / real_of_int j \ range of_rat" + by blast qed qed -lemma Rats_eq_int_div_nat: - "\ = { real_of_int i / real n |i n. n \ 0}" -proof(auto simp:Rats_eq_int_div_int) - fix i j::int assume "j \ 0" - show "EX (i'::int) (n::nat). real_of_int i / real_of_int j = real_of_int i'/real n \ 00" - hence "real_of_int i / real_of_int j = real_of_int i/real(nat j) \ 0 = { real_of_int i / real n | i n. n \ 0}" +proof (auto simp: Rats_eq_int_div_int) + fix i j :: int + assume "j \ 0" + show "\(i'::int) (n::nat). real_of_int i / real_of_int j = real_of_int i' / real n \ 0 < n" + proof (cases "j > 0") + case True + then have "real_of_int i / real_of_int j = real_of_int i / real (nat j) \ 0 < nat j" + by simp + then show ?thesis by blast next - assume "~ j>0" - hence "real_of_int i / real_of_int j = real_of_int(-i) / real(nat(-j)) \ 0j\0\ - by (simp add: of_nat_nat) - thus ?thesis by blast + case False + with \j \ 0\ + have "real_of_int i / real_of_int j = real_of_int (- i) / real (nat (- j)) \ 0 < nat (- j)" + by simp + then show ?thesis by blast qed next - fix i::int and n::nat assume "0 < n" - hence "real_of_int i / real n = real_of_int i / real_of_int(int n) \ int n \ 0" by simp - thus "\i' j. real_of_int i / real n = real_of_int i' / real_of_int j \ j \ 0" by blast + fix i :: int and n :: nat + assume "0 < n" + then have "real_of_int i / real n = real_of_int i / real_of_int(int n) \ int n \ 0" + by simp + then show "\i' j. real_of_int i / real n = real_of_int i' / real_of_int j \ j \ 0" + by blast qed lemma Rats_abs_nat_div_natE: assumes "x \ \" - obtains m n :: nat - where "n \ 0" and "\x\ = real m / real n" and "gcd m n = 1" + obtains m n :: nat where "n \ 0" and "\x\ = real m / real n" and "gcd m n = 1" proof - - from \x \ \\ obtain i::int and n::nat where "n \ 0" and "x = real_of_int i / real n" - by(auto simp add: Rats_eq_int_div_nat) - hence "\x\ = real (nat \i\) / real n" by (simp add: of_nat_nat) + from \x \ \\ obtain i :: int and n :: nat where "n \ 0" and "x = real_of_int i / real n" + by (auto simp add: Rats_eq_int_div_nat) + then have "\x\ = real (nat \i\) / real n" by simp then obtain m :: nat where x_rat: "\x\ = real m / real n" by blast let ?gcd = "gcd m n" - from \n\0\ have gcd: "?gcd \ 0" by simp + from \n \ 0\ have gcd: "?gcd \ 0" by simp let ?k = "m div ?gcd" let ?l = "n div ?gcd" let ?gcd' = "gcd ?k ?l" - have "?gcd dvd m" .. then have gcd_k: "?gcd * ?k = m" + have "?gcd dvd m" .. + then have gcd_k: "?gcd * ?k = m" by (rule dvd_mult_div_cancel) - have "?gcd dvd n" .. then have gcd_l: "?gcd * ?l = n" + have "?gcd dvd n" .. + then have gcd_l: "?gcd * ?l = n" by (rule dvd_mult_div_cancel) - from \n \ 0\ and gcd_l - have "?gcd * ?l \ 0" by simp + from \n \ 0\ and gcd_l have "?gcd * ?l \ 0" by simp then have "?l \ 0" by (blast dest!: mult_not_zero) moreover have "\x\ = real ?k / real ?l" @@ -1237,19 +1245,22 @@ ultimately show ?thesis .. qed -subsection\Density of the Rational Reals in the Reals\ + +subsection \Density of the Rational Reals in the Reals\ -text\This density proof is due to Stefan Richter and was ported by TN. The -original source is \emph{Real Analysis} by H.L. Royden. -It employs the Archimedean property of the reals.\ +text \ + This density proof is due to Stefan Richter and was ported by TN. The + original source is \emph{Real Analysis} by H.L. Royden. + It employs the Archimedean property of the reals.\ lemma Rats_dense_in_real: fixes x :: real - assumes "x < y" shows "\r\\. x < r \ r < y" + assumes "x < y" + shows "\r\\. x < r \ r < y" proof - - from \x have "0 < y-x" by simp - with reals_Archimedean obtain q::nat - where q: "inverse (real q) < y-x" and "0 < q" by blast + from \x < y\ have "0 < y - x" by simp + with reals_Archimedean obtain q :: nat where q: "inverse (real q) < y - x" and "0 < q" + by blast define p where "p = \y * real q\ - 1" define r where "r = of_int p / real q" from q have "x < y - inverse (real q)" by simp @@ -1259,8 +1270,7 @@ finally have "x < r" . moreover have "r < y" unfolding r_def p_def - by (simp add: divide_less_eq diff_less_eq \0 < q\ - less_ceiling_iff [symmetric]) + by (simp add: divide_less_eq diff_less_eq \0 < q\ less_ceiling_iff [symmetric]) moreover from r_def have "r \ \" by simp ultimately show ?thesis by blast qed @@ -1269,11 +1279,11 @@ fixes x y :: real assumes "x < y" shows "\q :: rat. x < of_rat q \ of_rat q < y" -using Rats_dense_in_real [OF \x < y\] -by (auto elim: Rats_cases) + using Rats_dense_in_real [OF \x < y\] + by (auto elim: Rats_cases) -subsection\Numerals and Arithmetic\ +subsection \Numerals and Arithmetic\ lemma [code_abbrev]: (*FIXME*) "real_of_int (numeral k) = numeral k" @@ -1294,147 +1304,140 @@ #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int \ real"})) \ -subsection\Simprules combining x+y and 0: ARE THEY NEEDED?\ + +subsection \Simprules combining \x + y\ and \0\\ (* FIXME ARE THEY NEEDED? *) -lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" -by arith +lemma real_add_minus_iff [simp]: "x + - a = 0 \ x = a" for x a :: real + by arith -lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)" -by auto +lemma real_add_less_0_iff: "x + y < 0 \ y < - x" for x y :: real + by auto -lemma real_0_less_add_iff: "((0::real) < x+y) = (-x < y)" -by auto +lemma real_0_less_add_iff: "0 < x + y \ - x < y" for x y :: real + by auto -lemma real_add_le_0_iff: "(x+y \ (0::real)) = (y \ -x)" -by auto +lemma real_add_le_0_iff: "x + y \ 0 \ y \ - x" for x y :: real + by auto -lemma real_0_le_add_iff: "((0::real) \ x+y) = (-x \ y)" -by auto +lemma real_0_le_add_iff: "0 \ x + y \ - x \ y" for x y :: real + by auto + subsection \Lemmas about powers\ lemma two_realpow_ge_one: "(1::real) \ 2 ^ n" by simp -text \FIXME: declare this [simp] for all types, or not at all\ +(* FIXME: declare this [simp] for all types, or not at all *) declare sum_squares_eq_zero_iff [simp] sum_power2_eq_zero_iff [simp] -lemma real_minus_mult_self_le [simp]: "-(u * u) \ (x * (x::real))" -by (rule_tac y = 0 in order_trans, auto) +lemma real_minus_mult_self_le [simp]: "- (u * u) \ x * x" for u x :: real + by (rule order_trans [where y = 0]) auto -lemma realpow_square_minus_le [simp]: "- u\<^sup>2 \ (x::real)\<^sup>2" +lemma realpow_square_minus_le [simp]: "- u\<^sup>2 \ x\<^sup>2" for u x :: real by (auto simp add: power2_eq_square) -lemma numeral_power_eq_real_of_int_cancel_iff[simp]: - "numeral x ^ n = real_of_int (y::int) \ numeral x ^ n = y" +lemma numeral_power_eq_real_of_int_cancel_iff [simp]: + "numeral x ^ n = real_of_int y \ numeral x ^ n = y" by (metis of_int_eq_iff of_int_numeral of_int_power) -lemma real_of_int_eq_numeral_power_cancel_iff[simp]: - "real_of_int (y::int) = numeral x ^ n \ y = numeral x ^ n" - using numeral_power_eq_real_of_int_cancel_iff[of x n y] - by metis +lemma real_of_int_eq_numeral_power_cancel_iff [simp]: + "real_of_int y = numeral x ^ n \ y = numeral x ^ n" + using numeral_power_eq_real_of_int_cancel_iff [of x n y] by metis -lemma numeral_power_eq_real_of_nat_cancel_iff[simp]: - "numeral x ^ n = real (y::nat) \ numeral x ^ n = y" +lemma numeral_power_eq_real_of_nat_cancel_iff [simp]: + "numeral x ^ n = real y \ numeral x ^ n = y" using of_nat_eq_iff by fastforce -lemma real_of_nat_eq_numeral_power_cancel_iff[simp]: - "real (y::nat) = numeral x ^ n \ y = numeral x ^ n" - using numeral_power_eq_real_of_nat_cancel_iff[of x n y] - by metis +lemma real_of_nat_eq_numeral_power_cancel_iff [simp]: + "real y = numeral x ^ n \ y = numeral x ^ n" + using numeral_power_eq_real_of_nat_cancel_iff [of x n y] by metis -lemma numeral_power_le_real_of_nat_cancel_iff[simp]: - "(numeral x::real) ^ n \ real a \ (numeral x::nat) ^ n \ a" -by (metis of_nat_le_iff of_nat_numeral of_nat_power) +lemma numeral_power_le_real_of_nat_cancel_iff [simp]: + "(numeral x :: real) ^ n \ real a \ (numeral x::nat) ^ n \ a" + by (metis of_nat_le_iff of_nat_numeral of_nat_power) -lemma real_of_nat_le_numeral_power_cancel_iff[simp]: +lemma real_of_nat_le_numeral_power_cancel_iff [simp]: "real a \ (numeral x::real) ^ n \ a \ (numeral x::nat) ^ n" -by (metis of_nat_le_iff of_nat_numeral of_nat_power) + by (metis of_nat_le_iff of_nat_numeral of_nat_power) -lemma numeral_power_le_real_of_int_cancel_iff[simp]: - "(numeral x::real) ^ n \ real_of_int a \ (numeral x::int) ^ n \ a" +lemma numeral_power_le_real_of_int_cancel_iff [simp]: + "(numeral x::real) ^ n \ real_of_int a \ (numeral x::int) ^ n \ a" by (metis ceiling_le_iff ceiling_of_int of_int_numeral of_int_power) -lemma real_of_int_le_numeral_power_cancel_iff[simp]: - "real_of_int a \ (numeral x::real) ^ n \ a \ (numeral x::int) ^ n" +lemma real_of_int_le_numeral_power_cancel_iff [simp]: + "real_of_int a \ (numeral x::real) ^ n \ a \ (numeral x::int) ^ n" by (metis floor_of_int le_floor_iff of_int_numeral of_int_power) -lemma numeral_power_less_real_of_nat_cancel_iff[simp]: - "(numeral x::real) ^ n < real a \ (numeral x::nat) ^ n < a" +lemma numeral_power_less_real_of_nat_cancel_iff [simp]: + "(numeral x::real) ^ n < real a \ (numeral x::nat) ^ n < a" + by (metis of_nat_less_iff of_nat_numeral of_nat_power) + +lemma real_of_nat_less_numeral_power_cancel_iff [simp]: + "real a < (numeral x::real) ^ n \ a < (numeral x::nat) ^ n" by (metis of_nat_less_iff of_nat_numeral of_nat_power) -lemma real_of_nat_less_numeral_power_cancel_iff[simp]: - "real a < (numeral x::real) ^ n \ a < (numeral x::nat) ^ n" -by (metis of_nat_less_iff of_nat_numeral of_nat_power) - -lemma numeral_power_less_real_of_int_cancel_iff[simp]: - "(numeral x::real) ^ n < real_of_int a \ (numeral x::int) ^ n < a" +lemma numeral_power_less_real_of_int_cancel_iff [simp]: + "(numeral x::real) ^ n < real_of_int a \ (numeral x::int) ^ n < a" by (meson not_less real_of_int_le_numeral_power_cancel_iff) -lemma real_of_int_less_numeral_power_cancel_iff[simp]: - "real_of_int a < (numeral x::real) ^ n \ a < (numeral x::int) ^ n" +lemma real_of_int_less_numeral_power_cancel_iff [simp]: + "real_of_int a < (numeral x::real) ^ n \ a < (numeral x::int) ^ n" by (meson not_less numeral_power_le_real_of_int_cancel_iff) -lemma neg_numeral_power_le_real_of_int_cancel_iff[simp]: - "(- numeral x::real) ^ n \ real_of_int a \ (- numeral x::int) ^ n \ a" +lemma neg_numeral_power_le_real_of_int_cancel_iff [simp]: + "(- numeral x::real) ^ n \ real_of_int a \ (- numeral x::int) ^ n \ a" by (metis of_int_le_iff of_int_neg_numeral of_int_power) -lemma real_of_int_le_neg_numeral_power_cancel_iff[simp]: - "real_of_int a \ (- numeral x::real) ^ n \ a \ (- numeral x::int) ^ n" +lemma real_of_int_le_neg_numeral_power_cancel_iff [simp]: + "real_of_int a \ (- numeral x::real) ^ n \ a \ (- numeral x::int) ^ n" by (metis of_int_le_iff of_int_neg_numeral of_int_power) -subsection\Density of the Reals\ +subsection \Density of the Reals\ + +lemma real_lbound_gt_zero: "0 < d1 \ 0 < d2 \ \e. 0 < e \ e < d1 \ e < d2" for d1 d2 :: real + by (rule exI [where x = "min d1 d2 / 2"]) (simp add: min_def) -lemma real_lbound_gt_zero: - "[| (0::real) < d1; 0 < d2 |] ==> \e. 0 < e & e < d1 & e < d2" -apply (rule_tac x = " (min d1 d2) /2" in exI) -apply (simp add: min_def) -done +text \Similar results are proved in @{theory Fields}\ +lemma real_less_half_sum: "x < y \ x < (x + y) / 2" for x y :: real + by auto + +lemma real_gt_half_sum: "x < y \ (x + y) / 2 < y" for x y :: real + by auto + +lemma real_sum_of_halves: "x / 2 + x / 2 = x" for x :: real + by simp -text\Similar results are proved in \Fields\\ -lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)" - by auto - -lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y" - by auto - -lemma real_sum_of_halves: "x/2 + x/2 = (x::real)" - by simp - -subsection\Floor and Ceiling Functions from the Reals to the Integers\ +subsection \Floor and Ceiling Functions from the Reals to the Integers\ (* FIXME: theorems for negative numerals. Many duplicates, e.g. from Archimedean_Field.thy. *) -lemma real_of_nat_less_numeral_iff [simp]: - "real (n::nat) < numeral w \ n < numeral w" +lemma real_of_nat_less_numeral_iff [simp]: "real n < numeral w \ n < numeral w" for n :: nat by (metis of_nat_less_iff of_nat_numeral) -lemma numeral_less_real_of_nat_iff [simp]: - "numeral w < real (n::nat) \ numeral w < n" +lemma numeral_less_real_of_nat_iff [simp]: "numeral w < real n \ numeral w < n" for n :: nat by (metis of_nat_less_iff of_nat_numeral) -lemma numeral_le_real_of_nat_iff[simp]: - "(numeral n \ real(m::nat)) = (numeral n \ m)" -by (metis not_le real_of_nat_less_numeral_iff) +lemma numeral_le_real_of_nat_iff [simp]: "numeral n \ real m \ numeral n \ m" for m :: nat + by (metis not_le real_of_nat_less_numeral_iff) -declare of_int_floor_le [simp] (* FIXME*) +declare of_int_floor_le [simp] (* FIXME duplicate!? *) -lemma of_int_floor_cancel [simp]: - "(of_int \x\ = x) = (\n::int. x = of_int n)" +lemma of_int_floor_cancel [simp]: "of_int \x\ = x \ (\n::int. x = of_int n)" by (metis floor_of_int) -lemma floor_eq: "[| real_of_int n < x; x < real_of_int n + 1 |] ==> \x\ = n" +lemma floor_eq: "real_of_int n < x \ x < real_of_int n + 1 \ \x\ = n" by linarith -lemma floor_eq2: "[| real_of_int n \ x; x < real_of_int n + 1 |] ==> \x\ = n" +lemma floor_eq2: "real_of_int n \ x \ x < real_of_int n + 1 \ \x\ = n" by linarith -lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat \x\ = n" +lemma floor_eq3: "real n < x \ x < real (Suc n) \ nat \x\ = n" by linarith -lemma floor_eq4: "[| real n \ x; x < real (Suc n) |] ==> nat \x\ = n" +lemma floor_eq4: "real n \ x \ x < real (Suc n) \ nat \x\ = n" by linarith lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real_of_int \r\" @@ -1455,41 +1458,52 @@ lemma floor_add2[simp]: "\of_int a + x\ = a + \x\" by (simp add: add.commute) -lemma floor_divide_real_eq_div: "0 \ b \ \a / real_of_int b\ = \a\ div b" -proof cases - assume "0 < b" - { fix i j :: int assume "real_of_int i \ a" "a < 1 + real_of_int i" +lemma floor_divide_real_eq_div: + assumes "0 \ b" + shows "\a / real_of_int b\ = \a\ div b" +proof (cases "b = 0") + case True + then show ?thesis by simp +next + case False + with assms have b: "b > 0" by simp + have "j = i div b" + if "real_of_int i \ a" "a < 1 + real_of_int i" "real_of_int j * real_of_int b \ a" "a < real_of_int b + real_of_int j * real_of_int b" - then have "i < b + j * b" - by (metis linorder_not_less of_int_add of_int_le_iff of_int_mult order_trans_rules(21)) + for i j :: int + proof - + from that have "i < b + j * b" + by (metis le_less_trans of_int_add of_int_less_iff of_int_mult) moreover have "j * b < 1 + i" proof - have "real_of_int (j * b) < real_of_int i + 1" using \a < 1 + real_of_int i\ \real_of_int j * real_of_int b \ a\ by force - thus "j * b < 1 + i" + then show "j * b < 1 + i" by linarith qed ultimately have "(j - i div b) * b \ i mod b" "i mod b < ((j - i div b) + 1) * b" by (auto simp: field_simps) then have "(j - i div b) * b < 1 * b" "0 * b < ((j - i div b) + 1) * b" - using pos_mod_bound[OF \0, of i] pos_mod_sign[OF \0, of i] by linarith+ - then have "j = i div b" - using \0 < b\ unfolding mult_less_cancel_right by auto - } - with \0 < b\ show ?thesis + using pos_mod_bound [OF b, of i] pos_mod_sign [OF b, of i] + by linarith+ + then show ?thesis + using b unfolding mult_less_cancel_right by auto + qed + with b show ?thesis by (auto split: floor_split simp: field_simps) -qed auto +qed -lemma floor_divide_eq_div_numeral[simp]: "\numeral a / numeral b::real\ = numeral a div numeral b" +lemma floor_divide_eq_div_numeral [simp]: "\numeral a / numeral b::real\ = numeral a div numeral b" by (metis floor_divide_of_int_eq of_int_numeral) -lemma floor_minus_divide_eq_div_numeral[simp]: "\- (numeral a / numeral b)::real\ = - numeral a div numeral b" +lemma floor_minus_divide_eq_div_numeral [simp]: + "\- (numeral a / numeral b)::real\ = - numeral a div numeral b" by (metis divide_minus_left floor_divide_of_int_eq of_int_neg_numeral of_int_numeral) -lemma of_int_ceiling_cancel [simp]: "(of_int \x\ = x) = (\n::int. x = of_int n)" +lemma of_int_ceiling_cancel [simp]: "of_int \x\ = x \ (\n::int. x = of_int n)" using ceiling_of_int by metis -lemma ceiling_eq: "[| of_int n < x; x \ of_int n + 1 |] ==> \x\ = n + 1" +lemma ceiling_eq: "of_int n < x \ x \ of_int n + 1 \ \x\ = n + 1" by (simp add: ceiling_unique) lemma of_int_ceiling_diff_one_le [simp]: "of_int \r\ - 1 \ r" @@ -1498,7 +1512,7 @@ lemma of_int_ceiling_le_add_one [simp]: "of_int \r\ \ r + 1" by linarith -lemma ceiling_le: "x <= of_int a ==> \x\ <= a" +lemma ceiling_le: "x \ of_int a \ \x\ \ a" by (simp add: ceiling_le_iff) lemma ceiling_divide_eq_div: "\of_int a / of_int b\ = - (- a div b)" @@ -1512,34 +1526,37 @@ "\- (numeral a / numeral b :: real)\ = - (numeral a div numeral b)" using ceiling_divide_eq_div[of "- numeral a" "numeral b"] by simp -text\The following lemmas are remnants of the erstwhile functions natfloor -and natceiling.\ +text \ + The following lemmas are remnants of the erstwhile functions natfloor + and natceiling. +\ -lemma nat_floor_neg: "(x::real) <= 0 ==> nat \x\ = 0" +lemma nat_floor_neg: "x \ 0 \ nat \x\ = 0" for x :: real by linarith -lemma le_nat_floor: "real x <= a ==> x <= nat \a\" +lemma le_nat_floor: "real x \ a \ x \ nat \a\" by linarith lemma le_mult_nat_floor: "nat \a\ * nat \b\ \ nat \a * b\" - by (cases "0 <= a & 0 <= b") + by (cases "0 \ a \ 0 \ b") (auto simp add: nat_mult_distrib[symmetric] nat_mono le_mult_floor) -lemma nat_ceiling_le_eq [simp]: "(nat \x\ <= a) = (x <= real a)" +lemma nat_ceiling_le_eq [simp]: "nat \x\ \ a \ x \ real a" by linarith -lemma real_nat_ceiling_ge: "x <= real (nat \x\)" +lemma real_nat_ceiling_ge: "x \ real (nat \x\)" by linarith -lemma Rats_no_top_le: "\ q \ \. (x :: real) \ q" +lemma Rats_no_top_le: "\q \ \. x \ q" for x :: real by (auto intro!: bexI[of _ "of_nat (nat \x\)"]) linarith -lemma Rats_no_bot_less: "\ q \ \. q < (x :: real)" +lemma Rats_no_bot_less: "\q \ \. q < x" for x :: real apply (auto intro!: bexI[of _ "of_int (\x\ - 1)"]) apply (rule less_le_trans[OF _ of_int_floor_le]) apply simp done + subsection \Exponentiation with floor\ lemma floor_power: @@ -1551,48 +1568,41 @@ then show ?thesis by (metis floor_of_int) qed -lemma floor_numeral_power[simp]: - "\numeral x ^ n\ = numeral x ^ n" +lemma floor_numeral_power [simp]: "\numeral x ^ n\ = numeral x ^ n" by (metis floor_of_int of_int_numeral of_int_power) -lemma ceiling_numeral_power[simp]: - "\numeral x ^ n\ = numeral x ^ n" +lemma ceiling_numeral_power [simp]: "\numeral x ^ n\ = numeral x ^ n" by (metis ceiling_of_int of_int_numeral of_int_power) + subsection \Implementation of rational real numbers\ text \Formal constructor\ -definition Ratreal :: "rat \ real" where - [code_abbrev, simp]: "Ratreal = of_rat" +definition Ratreal :: "rat \ real" + where [code_abbrev, simp]: "Ratreal = of_rat" code_datatype Ratreal text \Numerals\ -lemma [code_abbrev]: - "(of_rat (of_int a) :: real) = of_int a" +lemma [code_abbrev]: "(of_rat (of_int a) :: real) = of_int a" by simp -lemma [code_abbrev]: - "(of_rat 0 :: real) = 0" +lemma [code_abbrev]: "(of_rat 0 :: real) = 0" by simp -lemma [code_abbrev]: - "(of_rat 1 :: real) = 1" +lemma [code_abbrev]: "(of_rat 1 :: real) = 1" by simp -lemma [code_abbrev]: - "(of_rat (- 1) :: real) = - 1" +lemma [code_abbrev]: "(of_rat (- 1) :: real) = - 1" by simp -lemma [code_abbrev]: - "(of_rat (numeral k) :: real) = numeral k" +lemma [code_abbrev]: "(of_rat (numeral k) :: real) = numeral k" by simp -lemma [code_abbrev]: - "(of_rat (- numeral k) :: real) = - numeral k" +lemma [code_abbrev]: "(of_rat (- numeral k) :: real) = - numeral k" by simp lemma [code_post]: @@ -1605,28 +1615,23 @@ text \Operations\ -lemma zero_real_code [code]: - "0 = Ratreal 0" +lemma zero_real_code [code]: "0 = Ratreal 0" by simp -lemma one_real_code [code]: - "1 = Ratreal 1" +lemma one_real_code [code]: "1 = Ratreal 1" by simp instantiation real :: equal begin -definition "HOL.equal (x::real) y \ x - y = 0" +definition "HOL.equal x y \ x - y = 0" for x :: real -instance proof -qed (simp add: equal_real_def) +instance by standard (simp add: equal_real_def) -lemma real_equal_code [code]: - "HOL.equal (Ratreal x) (Ratreal y) \ HOL.equal x y" +lemma real_equal_code [code]: "HOL.equal (Ratreal x) (Ratreal y) \ HOL.equal x y" by (simp add: equal_real_def equal) -lemma [code nbe]: - "HOL.equal (x::real) x \ True" +lemma [code nbe]: "HOL.equal x x \ True" for x :: real by (rule equal_refl) end @@ -1656,14 +1661,15 @@ by (simp add: of_rat_divide) lemma real_floor_code [code]: "\Ratreal x\ = \x\" - 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) + 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) text \Quickcheck\ definition (in term_syntax) - valterm_ratreal :: "rat \ (unit \ Code_Evaluation.term) \ real \ (unit \ Code_Evaluation.term)" where - [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\} k" + valterm_ratreal :: "rat \ (unit \ Code_Evaluation.term) \ real \ (unit \ Code_Evaluation.term)" + where [code_unfold]: "valterm_ratreal k = Code_Evaluation.valtermify Ratreal {\} k" notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) @@ -1685,7 +1691,7 @@ begin definition - "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (%r. f (Ratreal r)) d" + "exhaustive_real f d = Quickcheck_Exhaustive.exhaustive (\r. f (Ratreal r)) d" instance .. @@ -1695,7 +1701,7 @@ begin definition - "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (%r. f (valterm_ratreal r)) d" + "full_exhaustive_real f d = Quickcheck_Exhaustive.full_exhaustive (\r. f (valterm_ratreal r)) d" instance .. @@ -1705,7 +1711,7 @@ begin definition - "narrowing = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing" + "narrowing_real = Quickcheck_Narrowing.apply (Quickcheck_Narrowing.cons Ratreal) narrowing" instance .. @@ -1727,9 +1733,9 @@ \ lemmas [nitpick_unfold] = inverse_real_inst.inverse_real one_real_inst.one_real - ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real - times_real_inst.times_real uminus_real_inst.uminus_real - zero_real_inst.zero_real + ord_real_inst.less_real ord_real_inst.less_eq_real plus_real_inst.plus_real + times_real_inst.times_real uminus_real_inst.uminus_real + zero_real_inst.zero_real subsection \Setup for SMT\ @@ -1738,11 +1744,12 @@ ML_file "Tools/SMT/z3_real.ML" lemma [z3_rule]: - "0 + (x::real) = x" + "0 + x = x" "x + 0 = x" "0 * x = 0" "1 * x = x" "x + y = y + x" + for x y :: real by auto end