# HG changeset patch # User paulson # Date 1447165121 0 # Node ID 77b453bd616f24dd2dabe5eebebf144e672bc435 # Parent 933eb9e6a1cc19ea1ad17c4f07a716b3a6d8f1ef Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed. diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Complex.thy --- a/src/HOL/Complex.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Complex.thy Tue Nov 10 14:18:41 2015 +0000 @@ -166,7 +166,7 @@ lemma Im_divide_numeral [simp]: "Im (z / numeral w) = Im z / numeral w" by (simp add: Im_divide sqr_conv_mult) - + lemma Re_divide_of_nat: "Re (z / of_nat n) = Re z / of_nat n" by (cases n) (simp_all add: Re_divide divide_simps power2_eq_square del: of_nat_Suc) @@ -688,7 +688,7 @@ by (simp add: complex_eq_iff cos_add sin_add) lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" - by (induct n, simp_all add: real_of_nat_Suc algebra_simps cis_mult) + by (induct n, simp_all add: of_nat_Suc algebra_simps cis_mult) lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)" by (simp add: complex_eq_iff) @@ -757,8 +757,7 @@ have "\ ^ n = fact n *\<^sub>R (cos_coeff n + \ * sin_coeff n)" by (induct n) (simp_all add: sin_coeff_Suc cos_coeff_Suc complex_eq_iff Re_divide Im_divide field_simps - power2_eq_square real_of_nat_Suc add_nonneg_eq_0_iff - real_of_nat_def[symmetric]) + power2_eq_square of_nat_Suc add_nonneg_eq_0_iff) then have "(\ * complex_of_real b) ^ n /\<^sub>R fact n = of_real (cos_coeff n * b^n) + \ * of_real (sin_coeff n * b^n)" by (simp add: field_simps) } diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Decision_Procs/Approximation.thy Tue Nov 10 14:18:41 2015 +0000 @@ -51,7 +51,7 @@ lemma horner_bounds': fixes lb :: "nat \ nat \ nat \ float \ float" and ub :: "nat \ nat \ nat \ float \ float" - assumes "0 \ real x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" + assumes "0 \ real_of_float x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) @@ -69,7 +69,7 @@ next case (Suc n) thus ?case using lapprox_rat[of prec 1 "f j'"] using rapprox_rat[of 1 "f j'" prec] - Suc[where j'="Suc j'"] \0 \ real x\ + Suc[where j'="Suc j'"] \0 \ real_of_float x\ by (auto intro!: add_mono mult_left_mono float_round_down_le float_round_up_le order_trans[OF add_mono[OF _ float_plus_down_le]] order_trans[OF _ add_mono[OF _ float_plus_up_le]] @@ -87,7 +87,7 @@ lemma horner_bounds: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" - assumes "0 \ real x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" + assumes "0 \ real_of_float x" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) @@ -102,14 +102,14 @@ (is "?ub") proof - have "?lb \ ?ub" - using horner_bounds'[where lb=lb, OF \0 \ real x\ f_Suc lb_0 lb_Suc ub_0 ub_Suc] - unfolding horner_schema[where f=f, OF f_Suc] . + using horner_bounds'[where lb=lb, OF \0 \ real_of_float x\ f_Suc lb_0 lb_Suc ub_0 ub_Suc] + unfolding horner_schema[where f=f, OF f_Suc] by simp thus "?lb" and "?ub" by auto qed lemma horner_bounds_nonpos: fixes F :: "nat \ nat" and G :: "nat \ nat \ nat" - assumes "real x \ 0" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" + assumes "real_of_float x \ 0" and f_Suc: "\n. f (Suc n) = G ((F ^^ n) s) (f n)" and lb_0: "\ i k x. lb 0 i k x = 0" and lb_Suc: "\ n i k x. lb (Suc n) i k x = float_plus_down prec (lapprox_rat prec 1 k) @@ -118,14 +118,14 @@ and ub_Suc: "\ n i k x. ub (Suc n) i k x = float_plus_up prec (rapprox_rat prec 1 k) (float_round_up prec (x * (lb n (F i) (G i k) x)))" - shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") + shows "(lb n ((F ^^ j') s) (f j') x) \ (\j=0..j=0.. (ub n ((F ^^ j') s) (f j') x)" (is "?ub") proof - have diff_mult_minus: "x - y * z = x + - y * z" for x y z :: float by simp - have sum_eq: "(\j=0..j = 0..j=0..j = 0.. real (-x)" using assms by auto + have "0 \ real_of_float (-x)" using assms by auto from horner_bounds[where G=G and F=F and f=f and s=s and prec=prec and lb="\ n i k x. lb n i k (-x)" and ub="\ n i k x. ub n i k (-x)", unfolded lb_Suc ub_Suc diff_mult_minus, @@ -238,7 +238,7 @@ qed lemma sqrt_iteration_bound: - assumes "0 < real x" + assumes "0 < real_of_float x" shows "sqrt x < sqrt_iteration prec n x" proof (induct n) case 0 @@ -260,7 +260,7 @@ proof (rule mult_strict_right_mono, auto) show "m < 2^nat (bitlen m)" using bitlen_bounds[OF \0 < m\, THEN conjunct2] - unfolding real_of_int_less_iff[of m, symmetric] by auto + unfolding of_int_less_iff[of m, symmetric] by auto qed finally have "sqrt x < sqrt (2 powr (e + bitlen m))" unfolding int_nat_bl by auto @@ -287,7 +287,7 @@ have E_eq: "2 powr ?E = 2 powr (?E div 2 + ?E div 2 + ?E mod 2)" by auto have "sqrt (2 powr ?E) = sqrt (2 powr (?E div 2) * 2 powr (?E div 2) * 2 powr (?E mod 2))" - unfolding E_eq unfolding powr_add[symmetric] by (simp add: int_of_reals del: real_of_ints) + unfolding E_eq unfolding powr_add[symmetric] by (metis of_int_add) also have "\ = 2 powr (?E div 2) * sqrt (2 powr (?E mod 2))" unfolding real_sqrt_mult[of _ "2 powr (?E mod 2)"] real_sqrt_abs2 by auto also have "\ < 2 powr (?E div 2) * 2 powr 1" @@ -304,11 +304,11 @@ case (Suc n) let ?b = "sqrt_iteration prec n x" have "0 < sqrt x" - using \0 < real x\ by auto - also have "\ < real ?b" + using \0 < real_of_float x\ by auto + also have "\ < real_of_float ?b" using Suc . finally have "sqrt x < (?b + x / ?b)/2" - using sqrt_ub_pos_pos_1[OF Suc _ \0 < real x\] by auto + using sqrt_ub_pos_pos_1[OF Suc _ \0 < real_of_float x\] by auto also have "\ \ (?b + (float_divr prec x ?b))/2" by (rule divide_right_mono, auto simp add: float_divr) also have "\ = (Float 1 (- 1)) * (?b + (float_divr prec x ?b))" @@ -320,8 +320,8 @@ qed lemma sqrt_iteration_lower_bound: - assumes "0 < real x" - shows "0 < real (sqrt_iteration prec n x)" (is "0 < ?sqrt") + assumes "0 < real_of_float x" + shows "0 < real_of_float (sqrt_iteration prec n x)" (is "0 < ?sqrt") proof - have "0 < sqrt x" using assms by auto also have "\ < ?sqrt" using sqrt_iteration_bound[OF assms] . @@ -329,21 +329,21 @@ qed lemma lb_sqrt_lower_bound: - assumes "0 \ real x" - shows "0 \ real (lb_sqrt prec x)" + assumes "0 \ real_of_float x" + shows "0 \ real_of_float (lb_sqrt prec x)" proof (cases "0 < x") case True - hence "0 < real x" and "0 \ x" - using \0 \ real x\ by auto + hence "0 < real_of_float x" and "0 \ x" + using \0 \ real_of_float x\ by auto hence "0 < sqrt_iteration prec prec x" using sqrt_iteration_lower_bound by auto - hence "0 \ real (float_divl prec x (sqrt_iteration prec prec x))" + hence "0 \ real_of_float (float_divl prec x (sqrt_iteration prec prec x))" using float_divl_lower_bound[OF \0 \ x\] unfolding less_eq_float_def by auto thus ?thesis unfolding lb_sqrt.simps using True by auto next case False - with \0 \ real x\ have "real x = 0" by auto + with \0 \ real_of_float x\ have "real_of_float x = 0" by auto thus ?thesis unfolding lb_sqrt.simps by auto qed @@ -352,24 +352,24 @@ proof - have lb: "lb_sqrt prec x \ sqrt x" if "0 < x" for x :: float proof - - from that have "0 < real x" and "0 \ real x" by auto + from that have "0 < real_of_float x" and "0 \ real_of_float x" by auto hence sqrt_gt0: "0 < sqrt x" by auto hence sqrt_ub: "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto have "(float_divl prec x (sqrt_iteration prec prec x)) \ x / (sqrt_iteration prec prec x)" by (rule float_divl) also have "\ < x / sqrt x" - by (rule divide_strict_left_mono[OF sqrt_ub \0 < real x\ + by (rule divide_strict_left_mono[OF sqrt_ub \0 < real_of_float x\ mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]]) also have "\ = sqrt x" unfolding inverse_eq_iff_eq[of _ "sqrt x", symmetric] - sqrt_divide_self_eq[OF \0 \ real x\, symmetric] by auto + sqrt_divide_self_eq[OF \0 \ real_of_float x\, symmetric] by auto finally show ?thesis unfolding lb_sqrt.simps if_P[OF \0 < x\] by auto qed have ub: "sqrt x \ ub_sqrt prec x" if "0 < x" for x :: float proof - - from that have "0 < real x" by auto + from that have "0 < real_of_float x" by auto hence "0 < sqrt x" by auto hence "sqrt x < sqrt_iteration prec prec x" using sqrt_iteration_bound by auto @@ -419,7 +419,7 @@ (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub_arctan_horner prec n (k + 2) x)))" lemma arctan_0_1_bounds': - assumes "0 \ real y" "real y \ 1" + assumes "0 \ real_of_float y" "real_of_float y \ 1" and "even n" shows "arctan (sqrt y) \ {(sqrt y * lb_arctan_horner prec n 1 y) .. (sqrt y * ub_arctan_horner prec (Suc n) 1 y)}" @@ -452,7 +452,7 @@ note bounds = horner_bounds[where s=1 and f="\i. 2 * i + 1" and j'=0 and lb="\n i k x. lb_arctan_horner prec n k x" and ub="\n i k x. ub_arctan_horner prec n k x", - OF \0 \ real y\ F lb_arctan_horner.simps ub_arctan_horner.simps] + OF \0 \ real_of_float y\ F lb_arctan_horner.simps ub_arctan_horner.simps] have "(sqrt y * lb_arctan_horner prec n 1 y) \ arctan (sqrt y)" proof - @@ -479,7 +479,7 @@ qed lemma arctan_0_1_bounds: - assumes "0 \ real y" "real y \ 1" + assumes "0 \ real_of_float y" "real_of_float y \ 1" shows "arctan (sqrt y) \ {(sqrt y * lb_arctan_horner prec (get_even n) 1 y) .. (sqrt y * ub_arctan_horner prec (get_odd n) 1 y)}" @@ -532,47 +532,37 @@ qed lemma arctan_0_1_bounds_le: - assumes "0 \ x" "x \ 1" "0 < real xl" "real xl \ x * x" "x * x \ real xu" "real xu \ 1" + assumes "0 \ x" "x \ 1" "0 < real_of_float xl" "real_of_float xl \ x * x" "x * x \ real_of_float xu" "real_of_float xu \ 1" shows "arctan x \ {x * lb_arctan_horner p1 (get_even n) 1 xu .. x * ub_arctan_horner p2 (get_odd n) 1 xl}" proof - - from assms have "real xl \ 1" "sqrt (real xl) \ x" "x \ sqrt (real xu)" "0 \ real xu" - "0 \ real xl" "0 < sqrt (real xl)" + from assms have "real_of_float xl \ 1" "sqrt (real_of_float xl) \ x" "x \ sqrt (real_of_float xu)" "0 \ real_of_float xu" + "0 \ real_of_float xl" "0 < sqrt (real_of_float xl)" by (auto intro!: real_le_rsqrt real_le_lsqrt simp: power2_eq_square) - from arctan_0_1_bounds[OF \0 \ real xu\ \real xu \ 1\] - have "sqrt (real xu) * real (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan (sqrt (real xu))" + from arctan_0_1_bounds[OF \0 \ real_of_float xu\ \real_of_float xu \ 1\] + have "sqrt (real_of_float xu) * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan (sqrt (real_of_float xu))" by simp from arctan_mult_le[OF \0 \ x\ \x \ sqrt _\ this] - have "x * real (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan x" . + have "x * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \ arctan x" . moreover - from arctan_0_1_bounds[OF \0 \ real xl\ \real xl \ 1\] - have "arctan (sqrt (real xl)) \ sqrt (real xl) * real (ub_arctan_horner p2 (get_odd n) 1 xl)" + from arctan_0_1_bounds[OF \0 \ real_of_float xl\ \real_of_float xl \ 1\] + have "arctan (sqrt (real_of_float xl)) \ sqrt (real_of_float xl) * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" by simp from arctan_le_mult[OF \0 < sqrt xl\ \sqrt xl \ x\ this] - have "arctan x \ x * real (ub_arctan_horner p2 (get_odd n) 1 xl)" . + have "arctan x \ x * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" . ultimately show ?thesis by simp qed -lemma mult_nonneg_le_one: - fixes a :: real - assumes "0 \ a" "0 \ b" "a \ 1" "b \ 1" - shows "a * b \ 1" -proof - - have "a * b \ 1 * 1" - by (intro mult_mono assms) simp_all - thus ?thesis by simp -qed - lemma arctan_0_1_bounds_round: - assumes "0 \ real x" "real x \ 1" + assumes "0 \ real_of_float x" "real_of_float x \ 1" shows "arctan x \ - {real x * lb_arctan_horner p1 (get_even n) 1 (float_round_up (Suc p2) (x * x)) .. - real x * ub_arctan_horner p3 (get_odd n) 1 (float_round_down (Suc p4) (x * x))}" + {real_of_float x * lb_arctan_horner p1 (get_even n) 1 (float_round_up (Suc p2) (x * x)) .. + real_of_float x * ub_arctan_horner p3 (get_odd n) 1 (float_round_down (Suc p4) (x * x))}" using assms apply (cases "x > 0") apply (intro arctan_0_1_bounds_le) apply (auto simp: float_round_down.rep_eq float_round_up.rep_eq - intro!: truncate_up_le1 mult_nonneg_le_one truncate_down_le truncate_up_le truncate_down_pos + intro!: truncate_up_le1 mult_le_one truncate_down_le truncate_up_le truncate_down_pos mult_pos_pos) done @@ -614,14 +604,14 @@ let ?kl = "float_round_down (Suc prec) (?k * ?k)" have "1 div k = 0" using div_pos_pos_trivial[OF _ \1 < k\] by auto - have "0 \ real ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: \0 \ k\) - have "real ?k \ 1" + have "0 \ real_of_float ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: \0 \ k\) + have "real_of_float ?k \ 1" by (auto simp add: \0 < k\ \1 \ k\ less_imp_le - intro!: mult_nonneg_le_one order_trans[OF _ rapprox_rat] rapprox_rat_le1) + intro!: mult_le_one order_trans[OF _ rapprox_rat] rapprox_rat_le1) have "1 / k \ ?k" using rapprox_rat[where x=1 and y=k] by auto hence "arctan (1 / k) \ arctan ?k" by (rule arctan_monotone') also have "\ \ (?k * ub_arctan_horner prec (get_odd n) 1 ?kl)" - using arctan_0_1_bounds_round[OF \0 \ real ?k\ \real ?k \ 1\] + using arctan_0_1_bounds_round[OF \0 \ real_of_float ?k\ \real_of_float ?k \ 1\] by auto finally have "arctan (1 / k) \ ?k * ub_arctan_horner prec (get_odd n) 1 ?kl" . } note ub_arctan = this @@ -634,16 +624,16 @@ let ?ku = "float_round_up (Suc prec) (?k * ?k)" have "1 div k = 0" using div_pos_pos_trivial[OF _ \1 < k\] by auto have "1 / k \ 1" using \1 < k\ by auto - have "0 \ real ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one \0 \ k\] + have "0 \ real_of_float ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one \0 \ k\] by (auto simp add: \1 div k = 0\) - have "0 \ real (?k * ?k)" by simp - have "real ?k \ 1" using lapprox_rat by (rule order_trans, auto simp add: \1 / k \ 1\) - hence "real (?k * ?k) \ 1" using \0 \ real ?k\ by (auto intro!: mult_nonneg_le_one) + have "0 \ real_of_float (?k * ?k)" by simp + have "real_of_float ?k \ 1" using lapprox_rat by (rule order_trans, auto simp add: \1 / k \ 1\) + hence "real_of_float (?k * ?k) \ 1" using \0 \ real_of_float ?k\ by (auto intro!: mult_le_one) have "?k \ 1 / k" using lapprox_rat[where x=1 and y=k] by auto have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \ arctan ?k" - using arctan_0_1_bounds_round[OF \0 \ real ?k\ \real ?k \ 1\] + using arctan_0_1_bounds_round[OF \0 \ real_of_float ?k\ \real_of_float ?k \ 1\] by auto also have "\ \ arctan (1 / k)" using \?k \ 1 / k\ by (rule arctan_monotone') finally have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \ arctan (1 / k)" . @@ -711,11 +701,11 @@ declare lb_arctan_horner.simps[simp del] lemma lb_arctan_bound': - assumes "0 \ real x" + assumes "0 \ real_of_float x" shows "lb_arctan prec x \ arctan x" proof - have "\ x < 0" and "0 \ x" - using \0 \ real x\ by (auto intro!: truncate_up_le ) + using \0 \ real_of_float x\ by (auto intro!: truncate_up_le ) let "?ub_horner x" = "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))" @@ -725,15 +715,15 @@ show ?thesis proof (cases "x \ Float 1 (- 1)") case True - hence "real x \ 1" by simp - from arctan_0_1_bounds_round[OF \0 \ real x\ \real x \ 1\] + hence "real_of_float x \ 1" by simp + from arctan_0_1_bounds_round[OF \0 \ real_of_float x\ \real_of_float x \ 1\] show ?thesis unfolding lb_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_P[OF True] using \0 \ x\ by (auto intro!: float_round_down_le) next case False - hence "0 < real x" by auto - let ?R = "1 + sqrt (1 + real x * real x)" + hence "0 < real_of_float x" by auto + let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" let ?sxx = "float_plus_up prec 1 (float_round_up prec (x * x))" let ?fR = "float_plus_up prec 1 (ub_sqrt prec ?sxx)" let ?DIV = "float_divl prec x ?fR" @@ -747,12 +737,12 @@ finally have "sqrt (1 + x*x) \ ub_sqrt prec ?sxx" . hence "?R \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) - hence "0 < ?fR" and "0 < real ?fR" using \0 < ?R\ by auto + hence "0 < ?fR" and "0 < real_of_float ?fR" using \0 < ?R\ by auto have monotone: "?DIV \ x / ?R" proof - - have "?DIV \ real x / ?fR" by (rule float_divl) - also have "\ \ x / ?R" by (rule divide_left_mono[OF \?R \ ?fR\ \0 \ real x\ mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \?R \ real ?fR\] divisor_gt0]]) + have "?DIV \ real_of_float x / ?fR" by (rule float_divl) + also have "\ \ x / ?R" by (rule divide_left_mono[OF \?R \ ?fR\ \0 \ real_of_float x\ mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \?R \ real_of_float ?fR\] divisor_gt0]]) finally show ?thesis . qed @@ -762,18 +752,18 @@ have "x \ sqrt (1 + x * x)" using real_sqrt_sum_squares_ge2[where x=1, unfolded numeral_2_eq_2] by auto also note \\ \ (ub_sqrt prec ?sxx)\ - finally have "real x \ ?fR" + finally have "real_of_float x \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) - moreover have "?DIV \ real x / ?fR" + moreover have "?DIV \ real_of_float x / ?fR" by (rule float_divl) - ultimately have "real ?DIV \ 1" - unfolding divide_le_eq_1_pos[OF \0 < real ?fR\, symmetric] by auto - - have "0 \ real ?DIV" + ultimately have "real_of_float ?DIV \ 1" + unfolding divide_le_eq_1_pos[OF \0 < real_of_float ?fR\, symmetric] by auto + + have "0 \ real_of_float ?DIV" using float_divl_lower_bound[OF \0 \ x\] \0 < ?fR\ unfolding less_eq_float_def by auto - from arctan_0_1_bounds_round[OF \0 \ real (?DIV)\ \real (?DIV) \ 1\] + from arctan_0_1_bounds_round[OF \0 \ real_of_float (?DIV)\ \real_of_float (?DIV) \ 1\] have "Float 1 1 * ?lb_horner ?DIV \ 2 * arctan ?DIV" by simp also have "\ \ 2 * arctan (x / ?R)" @@ -787,11 +777,11 @@ intro!: order_trans[OF mult_left_mono[OF truncate_down]]) next case False - hence "2 < real x" by auto - hence "1 \ real x" by auto + hence "2 < real_of_float x" by auto + hence "1 \ real_of_float x" by auto let "?invx" = "float_divr prec 1 x" - have "0 \ arctan x" using arctan_monotone'[OF \0 \ real x\] + have "0 \ arctan x" using arctan_monotone'[OF \0 \ real_of_float x\] using arctan_tan[of 0, unfolded tan_zero] by auto show ?thesis @@ -803,22 +793,22 @@ using \0 \ arctan x\ by auto next case False - hence "real ?invx \ 1" by auto - have "0 \ real ?invx" - by (rule order_trans[OF _ float_divr]) (auto simp add: \0 \ real x\) + hence "real_of_float ?invx \ 1" by auto + have "0 \ real_of_float ?invx" + by (rule order_trans[OF _ float_divr]) (auto simp add: \0 \ real_of_float x\) have "1 / x \ 0" and "0 < 1 / x" - using \0 < real x\ by auto + using \0 < real_of_float x\ by auto have "arctan (1 / x) \ arctan ?invx" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone', rule float_divr) also have "\ \ ?ub_horner ?invx" - using arctan_0_1_bounds_round[OF \0 \ real ?invx\ \real ?invx \ 1\] + using arctan_0_1_bounds_round[OF \0 \ real_of_float ?invx\ \real_of_float ?invx \ 1\] by (auto intro!: float_round_up_le) also note float_round_up finally have "pi / 2 - float_round_up prec (?ub_horner ?invx) \ arctan x" using \0 \ arctan x\ arctan_inverse[OF \1 / x \ 0\] - unfolding real_sgn_pos[OF \0 < 1 / real x\] le_diff_eq by auto + unfolding real_sgn_pos[OF \0 < 1 / real_of_float x\] le_diff_eq by auto moreover have "lb_pi prec * Float 1 (- 1) \ pi / 2" unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by simp @@ -833,11 +823,11 @@ qed lemma ub_arctan_bound': - assumes "0 \ real x" + assumes "0 \ real_of_float x" shows "arctan x \ ub_arctan prec x" proof - have "\ x < 0" and "0 \ x" - using \0 \ real x\ by auto + using \0 \ real_of_float x\ by auto let "?ub_horner x" = "float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x)))" @@ -847,22 +837,22 @@ show ?thesis proof (cases "x \ Float 1 (- 1)") case True - hence "real x \ 1" by auto + hence "real_of_float x \ 1" by auto show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] if_P[OF True] - using arctan_0_1_bounds_round[OF \0 \ real x\ \real x \ 1\] + using arctan_0_1_bounds_round[OF \0 \ real_of_float x\ \real_of_float x \ 1\] by (auto intro!: float_round_up_le) next case False - hence "0 < real x" by auto - let ?R = "1 + sqrt (1 + real x * real x)" + hence "0 < real_of_float x" by auto + let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" let ?sxx = "float_plus_down prec 1 (float_round_down prec (x * x))" let ?fR = "float_plus_down (Suc prec) 1 (lb_sqrt prec ?sxx)" let ?DIV = "float_divr prec x ?fR" - have sqr_ge0: "0 \ 1 + real x * real x" - using sum_power2_ge_zero[of 1 "real x", unfolded numeral_2_eq_2] by auto - hence "0 \ real (1 + x*x)" by auto + have sqr_ge0: "0 \ 1 + real_of_float x * real_of_float x" + using sum_power2_ge_zero[of 1 "real_of_float x", unfolded numeral_2_eq_2] by auto + hence "0 \ real_of_float (1 + x*x)" by auto hence divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) @@ -873,13 +863,13 @@ finally have "lb_sqrt prec ?sxx \ sqrt (1 + x*x)" . hence "?fR \ ?R" by (auto simp: float_plus_down.rep_eq plus_down_def truncate_down_le) - have "0 < real ?fR" + have "0 < real_of_float ?fR" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq intro!: truncate_down_ge1 lb_sqrt_lower_bound order_less_le_trans[OF zero_less_one] truncate_down_nonneg add_nonneg_nonneg) have monotone: "x / ?R \ (float_divr prec x ?fR)" proof - - from divide_left_mono[OF \?fR \ ?R\ \0 \ real x\ mult_pos_pos[OF divisor_gt0 \0 < real ?fR\]] + from divide_left_mono[OF \?fR \ ?R\ \0 \ real_of_float x\ mult_pos_pos[OF divisor_gt0 \0 < real_of_float ?fR\]] have "x / ?R \ x / ?fR" . also have "\ \ ?DIV" by (rule float_divr) finally show ?thesis . @@ -899,11 +889,11 @@ if_not_P[OF \\ x \ Float 1 (- 1)\] if_P[OF \x \ Float 1 1\] if_P[OF True] . next case False - hence "real ?DIV \ 1" by auto + hence "real_of_float ?DIV \ 1" by auto have "0 \ x / ?R" - using \0 \ real x\ \0 < ?R\ unfolding zero_le_divide_iff by auto - hence "0 \ real ?DIV" + using \0 \ real_of_float x\ \0 < ?R\ unfolding zero_le_divide_iff by auto + hence "0 \ real_of_float ?DIV" using monotone by (rule order_trans) have "arctan x = 2 * arctan (x / ?R)" @@ -911,7 +901,7 @@ also have "\ \ 2 * arctan (?DIV)" using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono) also have "\ \ (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num - using arctan_0_1_bounds_round[OF \0 \ real ?DIV\ \real ?DIV \ 1\] + using arctan_0_1_bounds_round[OF \0 \ real_of_float ?DIV\ \real_of_float ?DIV \ 1\] by (auto intro!: float_round_up_le) finally show ?thesis unfolding ub_arctan.simps Let_def if_not_P[OF \\ x < 0\] @@ -919,27 +909,27 @@ qed next case False - hence "2 < real x" by auto - hence "1 \ real x" by auto - hence "0 < real x" by auto + hence "2 < real_of_float x" by auto + hence "1 \ real_of_float x" by auto + hence "0 < real_of_float x" by auto hence "0 < x" by auto let "?invx" = "float_divl prec 1 x" have "0 \ arctan x" - using arctan_monotone'[OF \0 \ real x\] and arctan_tan[of 0, unfolded tan_zero] by auto - - have "real ?invx \ 1" + using arctan_monotone'[OF \0 \ real_of_float x\] and arctan_tan[of 0, unfolded tan_zero] by auto + + have "real_of_float ?invx \ 1" unfolding less_float_def by (rule order_trans[OF float_divl]) - (auto simp add: \1 \ real x\ divide_le_eq_1_pos[OF \0 < real x\]) - have "0 \ real ?invx" + (auto simp add: \1 \ real_of_float x\ divide_le_eq_1_pos[OF \0 < real_of_float x\]) + have "0 \ real_of_float ?invx" using \0 < x\ by (intro float_divl_lower_bound) auto have "1 / x \ 0" and "0 < 1 / x" - using \0 < real x\ by auto + using \0 < real_of_float x\ by auto have "(?lb_horner ?invx) \ arctan (?invx)" - using arctan_0_1_bounds_round[OF \0 \ real ?invx\ \real ?invx \ 1\] + using arctan_0_1_bounds_round[OF \0 \ real_of_float ?invx\ \real_of_float ?invx \ 1\] by (auto intro!: float_round_down_le) also have "\ \ arctan (1 / x)" unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone') (rule float_divl) @@ -962,17 +952,17 @@ lemma arctan_boundaries: "arctan x \ {(lb_arctan prec x) .. (ub_arctan prec x)}" proof (cases "0 \ x") case True - hence "0 \ real x" by auto + hence "0 \ real_of_float x" by auto show ?thesis - using ub_arctan_bound'[OF \0 \ real x\] lb_arctan_bound'[OF \0 \ real x\] + using ub_arctan_bound'[OF \0 \ real_of_float x\] lb_arctan_bound'[OF \0 \ real_of_float x\] unfolding atLeastAtMost_iff by auto next case False let ?mx = "-x" - from False have "x < 0" and "0 \ real ?mx" + from False have "x < 0" and "0 \ real_of_float ?mx" by auto hence bounds: "lb_arctan prec ?mx \ arctan ?mx \ arctan ?mx \ ub_arctan prec ?mx" - using ub_arctan_bound'[OF \0 \ real ?mx\] lb_arctan_bound'[OF \0 \ real ?mx\] by auto + using ub_arctan_bound'[OF \0 \ real_of_float ?mx\] lb_arctan_bound'[OF \0 \ real_of_float ?mx\] by auto show ?thesis unfolding minus_float.rep_eq arctan_minus lb_arctan.simps[where x=x] ub_arctan.simps[where x=x] Let_def if_P[OF \x < 0\] @@ -1027,7 +1017,7 @@ shows "(lb_sin_cos_aux prec n 1 1 (x * x)) \ (\ i=0.. i=0.. (ub_sin_cos_aux prec n 1 1 (x * x))" (is "?ub") proof - - have "0 \ real (x * x)" by auto + have "0 \ real_of_float (x * x)" by auto let "?f n" = "fact (2 * n) :: nat" have f_eq: "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 1 * (((\i. i + 2) ^^ n) 1 + 1)" for n proof - @@ -1035,9 +1025,9 @@ then show ?thesis by auto qed from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, - OF \0 \ real (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] + OF \0 \ real_of_float (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] show ?lb and ?ub - by (auto simp add: power_mult power2_eq_square[of "real x"]) + by (auto simp add: power_mult power2_eq_square[of "real_of_float x"]) qed lemma lb_sin_cos_aux_zero_le_one: "lb_sin_cos_aux prec n i j 0 \ 1" @@ -1048,13 +1038,13 @@ by (cases n) (auto intro!: float_plus_up_le order_trans[OF _ rapprox_rat]) lemma cos_boundaries: - assumes "0 \ real x" and "x \ pi / 2" + assumes "0 \ real_of_float x" and "x \ pi / 2" shows "cos x \ {(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) .. (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))}" -proof (cases "real x = 0") +proof (cases "real_of_float x = 0") case False - hence "real x \ 0" by auto - hence "0 < x" and "0 < real x" - using \0 \ real x\ by auto + hence "real_of_float x \ 0" by auto + hence "0 < x" and "0 < real_of_float x" + using \0 \ real_of_float x\ by auto have "0 < x * x" using \0 < x\ by simp @@ -1074,11 +1064,11 @@ { fix n :: nat assume "0 < n" hence "0 < 2 * n" by auto - obtain t where "0 < t" and "t < real x" and - cos_eq: "cos x = (\ i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * (real x) ^ i) - + (cos (t + 1/2 * (2 * n) * pi) / (fact (2*n))) * (real x)^(2*n)" + obtain t where "0 < t" and "t < real_of_float x" and + cos_eq: "cos x = (\ i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * (real_of_float x) ^ i) + + (cos (t + 1/2 * (2 * n) * pi) / (fact (2*n))) * (real_of_float x)^(2*n)" (is "_ = ?SUM + ?rest / ?fact * ?pow") - using Maclaurin_cos_expansion2[OF \0 < real x\ \0 < 2 * n\] + using Maclaurin_cos_expansion2[OF \0 < real_of_float x\ \0 < 2 * n\] unfolding cos_coeff_def atLeast0LessThan by auto have "cos t * (- 1) ^ n = cos t * cos (n * pi) + sin t * sin (n * pi)" by auto @@ -1086,12 +1076,12 @@ also have "\ = ?rest" by auto finally have "cos t * (- 1) ^ n = ?rest" . moreover - have "t \ pi / 2" using \t < real x\ and \x \ pi / 2\ by auto + have "t \ pi / 2" using \t < real_of_float x\ and \x \ pi / 2\ by auto hence "0 \ cos t" using \0 < t\ and cos_ge_zero by auto ultimately have even: "even n \ 0 \ ?rest" and odd: "odd n \ 0 \ - ?rest " by auto have "0 < ?fact" by auto - have "0 < ?pow" using \0 < real x\ by auto + have "0 < ?pow" using \0 < real_of_float x\ by auto { assume "even n" @@ -1131,7 +1121,7 @@ case False hence "get_even n = 0" by auto have "- (pi / 2) \ x" - by (rule order_trans[OF _ \0 < real x\[THEN less_imp_le]]) auto + by (rule order_trans[OF _ \0 < real_of_float x\[THEN less_imp_le]]) auto with \x \ pi / 2\ show ?thesis unfolding \get_even n = 0\ lb_sin_cos_aux.simps minus_float.rep_eq zero_float.rep_eq using cos_ge_zero by auto @@ -1147,13 +1137,13 @@ qed lemma sin_aux: - assumes "0 \ real x" + assumes "0 \ real_of_float x" shows "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ (\ i=0.. i=0.. (x * ub_sin_cos_aux prec n 2 1 (x * x))" (is "?ub") proof - - have "0 \ real (x * x)" by auto + have "0 \ real_of_float (x * x)" by auto let "?f n" = "fact (2 * n + 1) :: nat" have f_eq: "?f (Suc n) = ?f n * ((\i. i + 2) ^^ n) 2 * (((\i. i + 2) ^^ n) 2 + 1)" for n proof - @@ -1162,22 +1152,22 @@ unfolding F by auto qed from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, - OF \0 \ real (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] - show "?lb" and "?ub" using \0 \ real x\ + OF \0 \ real_of_float (x * x)\ f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] + show "?lb" and "?ub" using \0 \ real_of_float x\ unfolding power_add power_one_right mult.assoc[symmetric] setsum_left_distrib[symmetric] - unfolding mult.commute[where 'a=real] real_fact_nat - by (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real x"]) + unfolding mult.commute[where 'a=real] of_nat_fact + by (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real_of_float x"]) qed lemma sin_boundaries: - assumes "0 \ real x" + assumes "0 \ real_of_float x" and "x \ pi / 2" shows "sin x \ {(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) .. (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))}" -proof (cases "real x = 0") +proof (cases "real_of_float x = 0") case False - hence "real x \ 0" by auto - hence "0 < x" and "0 < real x" - using \0 \ real x\ by auto + hence "real_of_float x \ 0" by auto + hence "0 < x" and "0 < real_of_float x" + using \0 \ real_of_float x\ by auto have "0 < x * x" using \0 < x\ by simp @@ -1198,18 +1188,18 @@ { fix n :: nat assume "0 < n" hence "0 < 2 * n + 1" by auto - obtain t where "0 < t" and "t < real x" and - sin_eq: "sin x = (\ i = 0 ..< 2 * n + 1. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real x) ^ i) - + (sin (t + 1/2 * (2 * n + 1) * pi) / (fact (2*n + 1))) * (real x)^(2*n + 1)" + obtain t where "0 < t" and "t < real_of_float x" and + sin_eq: "sin x = (\ i = 0 ..< 2 * n + 1. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i) + + (sin (t + 1/2 * (2 * n + 1) * pi) / (fact (2*n + 1))) * (real_of_float x)^(2*n + 1)" (is "_ = ?SUM + ?rest / ?fact * ?pow") - using Maclaurin_sin_expansion3[OF \0 < 2 * n + 1\ \0 < real x\] + using Maclaurin_sin_expansion3[OF \0 < 2 * n + 1\ \0 < real_of_float x\] unfolding sin_coeff_def atLeast0LessThan by auto have "?rest = cos t * (- 1) ^ n" - unfolding sin_add cos_add real_of_nat_add distrib_right distrib_left by auto + unfolding sin_add cos_add of_nat_add distrib_right distrib_left by auto moreover have "t \ pi / 2" - using \t < real x\ and \x \ pi / 2\ by auto + using \t < real_of_float x\ and \x \ pi / 2\ by auto hence "0 \ cos t" using \0 < t\ and cos_ge_zero by auto ultimately have even: "even n \ 0 \ ?rest" and odd: "odd n \ 0 \ - ?rest" @@ -1218,13 +1208,13 @@ have "0 < ?fact" by (simp del: fact_Suc) have "0 < ?pow" - using \0 < real x\ by (rule zero_less_power) + using \0 < real_of_float x\ by (rule zero_less_power) { assume "even n" have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \ - (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real x) ^ i)" - using sin_aux[OF \0 \ real x\] unfolding setsum_morph[symmetric] by auto + (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" + using sin_aux[OF \0 \ real_of_float x\] unfolding setsum_morph[symmetric] by auto also have "\ \ ?SUM" by auto also have "\ \ sin x" proof - @@ -1244,10 +1234,10 @@ by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) thus ?thesis unfolding sin_eq by auto qed - also have "\ \ (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real x) ^ i)" + also have "\ \ (\ i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" by auto also have "\ \ (x * ub_sin_cos_aux prec n 2 1 (x * x))" - using sin_aux[OF \0 \ real x\] unfolding setsum_morph[symmetric] by auto + using sin_aux[OF \0 \ real_of_float x\] unfolding setsum_morph[symmetric] by auto finally have "sin x \ (x * ub_sin_cos_aux prec n 2 1 (x * x))" . } note ub = this and lb } note ub = this(1) and lb = this(2) @@ -1262,7 +1252,7 @@ next case False hence "get_even n = 0" by auto - with \x \ pi / 2\ \0 \ real x\ + with \x \ pi / 2\ \0 \ real_of_float x\ show ?thesis unfolding \get_even n = 0\ ub_sin_cos_aux.simps minus_float.rep_eq using sin_ge_zero by auto @@ -1275,13 +1265,13 @@ case True thus ?thesis unfolding \n = 0\ get_even_def get_odd_def - using \real x = 0\ lapprox_rat[where x="-1" and y=1] by auto + using \real_of_float x = 0\ lapprox_rat[where x="-1" and y=1] by auto next case False with not0_implies_Suc obtain m where "n = Suc m" by blast thus ?thesis unfolding \n = Suc m\ get_even_def get_odd_def - using \real x = 0\ rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] + using \real_of_float x = 0\ rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] by (cases "even (Suc m)") auto qed qed @@ -1306,7 +1296,7 @@ else half (half (horner (x * Float 1 (- 2)))))" lemma lb_cos: - assumes "0 \ real x" and "x \ pi" + assumes "0 \ real_of_float x" and "x \ pi" shows "cos x \ {(lb_cos prec x) .. (ub_cos prec x)}" (is "?cos x \ {(?lb x) .. (?ub x) }") proof - have x_half[symmetric]: "cos x = 2 * cos (x / 2) * cos (x / 2) - 1" for x :: real @@ -1320,7 +1310,7 @@ finally show ?thesis . qed - have "\ x < 0" using \0 \ real x\ by auto + have "\ x < 0" using \0 \ real_of_float x\ by auto let "?ub_horner x" = "ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x)" let "?lb_horner x" = "lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x)" let "?ub_half x" = "float_plus_up prec (Float 1 1 * x * x) (- 1)" @@ -1334,7 +1324,7 @@ show ?thesis unfolding lb_cos_def[where x=x] ub_cos_def[where x=x] if_not_P[OF \\ x < 0\] if_P[OF \x < Float 1 (- 1)\] Let_def - using cos_boundaries[OF \0 \ real x\ \x \ pi / 2\] . + using cos_boundaries[OF \0 \ real_of_float x\ \x \ pi / 2\] . next case False { fix y x :: float let ?x2 = "(x * Float 1 (- 1))" @@ -1351,12 +1341,12 @@ using cos_ge_minus_one unfolding if_P[OF True] by auto next case False - hence "0 \ real y" by auto + hence "0 \ real_of_float y" by auto from mult_mono[OF \y \ cos ?x2\ \y \ cos ?x2\ \0 \ cos ?x2\ this] - have "real y * real y \ cos ?x2 * cos ?x2" . - hence "2 * real y * real y \ 2 * cos ?x2 * cos ?x2" + have "real_of_float y * real_of_float y \ cos ?x2 * cos ?x2" . + hence "2 * real_of_float y * real_of_float y \ 2 * cos ?x2 * cos ?x2" by auto - hence "2 * real y * real y - 1 \ 2 * cos (x / 2) * cos (x / 2) - 1" + hence "2 * real_of_float y * real_of_float y - 1 \ 2 * cos (x / 2) * cos (x / 2) - 1" unfolding Float_num by auto thus ?thesis unfolding if_not_P[OF False] x_half Float_num @@ -1372,13 +1362,13 @@ have "cos x \ (?ub_half y)" proof - - have "0 \ real y" + have "0 \ real_of_float y" using \0 \ cos ?x2\ ub by (rule order_trans) from mult_mono[OF ub ub this \0 \ cos ?x2\] - have "cos ?x2 * cos ?x2 \ real y * real y" . - hence "2 * cos ?x2 * cos ?x2 \ 2 * real y * real y" + have "cos ?x2 * cos ?x2 \ real_of_float y * real_of_float y" . + hence "2 * cos ?x2 * cos ?x2 \ 2 * real_of_float y * real_of_float y" by auto - hence "2 * cos (x / 2) * cos (x / 2) - 1 \ 2 * real y * real y - 1" + hence "2 * cos (x / 2) * cos (x / 2) - 1 \ 2 * real_of_float y * real_of_float y - 1" unfolding Float_num by auto thus ?thesis unfolding x_half Float_num @@ -1390,15 +1380,15 @@ let ?x4 = "x * Float 1 (- 1) * Float 1 (- 1)" have "-pi \ x" - using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] \0 \ real x\ + using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] \0 \ real_of_float x\ by (rule order_trans) show ?thesis proof (cases "x < 1") case True - hence "real x \ 1" by auto - have "0 \ real ?x2" and "?x2 \ pi / 2" - using pi_ge_two \0 \ real x\ using assms by auto + hence "real_of_float x \ 1" by auto + have "0 \ real_of_float ?x2" and "?x2 \ pi / 2" + using pi_ge_two \0 \ real_of_float x\ using assms by auto from cos_boundaries[OF this] have lb: "(?lb_horner ?x2) \ ?cos ?x2" and ub: "?cos ?x2 \ (?ub_horner ?x2)" by auto @@ -1420,8 +1410,8 @@ ultimately show ?thesis by auto next case False - have "0 \ real ?x4" and "?x4 \ pi / 2" - using pi_ge_two \0 \ real x\ \x \ pi\ unfolding Float_num by auto + have "0 \ real_of_float ?x4" and "?x4 \ pi / 2" + using pi_ge_two \0 \ real_of_float x\ \x \ pi\ unfolding Float_num by auto from cos_boundaries[OF this] have lb: "(?lb_horner ?x4) \ ?cos ?x4" and ub: "?cos ?x4 \ (?ub_horner ?x4)" by auto @@ -1432,7 +1422,7 @@ have "(?lb x) \ ?cos x" proof - have "-pi \ ?x2" and "?x2 \ pi" - using pi_ge_two \0 \ real x\ \x \ pi\ by auto + using pi_ge_two \0 \ real_of_float x\ \x \ pi\ by auto from lb_half[OF lb_half[OF lb this] \-pi \ x\ \x \ pi\, unfolded eq_4] show ?thesis unfolding lb_cos_def[where x=x] if_not_P[OF \\ x < 0\] @@ -1441,7 +1431,7 @@ moreover have "?cos x \ (?ub x)" proof - have "-pi \ ?x2" and "?x2 \ pi" - using pi_ge_two \0 \ real x\ \ x \ pi\ by auto + using pi_ge_two \0 \ real_of_float x\ \ x \ pi\ by auto from ub_half[OF ub_half[OF ub this] \-pi \ x\ \x \ pi\, unfolded eq_4] show ?thesis unfolding ub_cos_def[where x=x] if_not_P[OF \\ x < 0\] @@ -1454,11 +1444,11 @@ lemma lb_cos_minus: assumes "-pi \ x" - and "real x \ 0" - shows "cos (real(-x)) \ {(lb_cos prec (-x)) .. (ub_cos prec (-x))}" + and "real_of_float x \ 0" + shows "cos (real_of_float(-x)) \ {(lb_cos prec (-x)) .. (ub_cos prec (-x))}" proof - - have "0 \ real (-x)" and "(-x) \ pi" - using \-pi \ x\ \real x \ 0\ by auto + have "0 \ real_of_float (-x)" and "(-x) \ pi" + using \-pi \ x\ \real_of_float x \ 0\ by auto from lb_cos[OF this] show ?thesis . qed @@ -1476,7 +1466,7 @@ else if -2 * lpi \ lx \ ux \ 0 then (Float (- 1) 0, max (ub_cos prec (lx + 2 * lpi)) (ub_cos prec (-ux))) else (Float (- 1) 0, Float 1 0))" -lemma floor_int: obtains k :: int where "real k = (floor_fl f)" +lemma floor_int: obtains k :: int where "real_of_int k = (floor_fl f)" by (simp add: floor_fl_def) lemma cos_periodic_nat[simp]: @@ -1488,7 +1478,7 @@ next case (Suc n) have split_pi_off: "x + (Suc n) * (2 * pi) = (x + n * (2 * pi)) + 2 * pi" - unfolding Suc_eq_plus1 real_of_nat_add real_of_one distrib_right by auto + unfolding Suc_eq_plus1 of_nat_add of_int_1 distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed @@ -1498,7 +1488,7 @@ shows "cos (x + i * (2 * pi)) = cos x" proof (cases "0 \ i") case True - hence i_nat: "real i = nat i" by auto + hence i_nat: "real_of_int i = nat i" by auto show ?thesis unfolding i_nat by auto next @@ -1526,7 +1516,7 @@ let ?lx = "float_plus_down prec lx ?lx2" let ?ux = "float_plus_up prec ux ?ux2" - obtain k :: int where k: "k = real ?k" + obtain k :: int where k: "k = real_of_float ?k" by (rule floor_int) have upi: "pi \ ?upi" and lpi: "?lpi \ pi" @@ -1542,18 +1532,18 @@ hence "?lx \ x - k * (2 * pi) \ x - k * (2 * pi) \ ?ux" by (auto intro!: float_plus_down_le float_plus_up_le) note lx = this[THEN conjunct1] and ux = this[THEN conjunct2] - hence lx_less_ux: "?lx \ real ?ux" by (rule order_trans) + hence lx_less_ux: "?lx \ real_of_float ?ux" by (rule order_trans) { assume "- ?lpi \ ?lx" and x_le_0: "x - k * (2 * pi) \ 0" with lpi[THEN le_imp_neg_le] lx - have pi_lx: "- pi \ ?lx" and lx_0: "real ?lx \ 0" + have pi_lx: "- pi \ ?lx" and lx_0: "real_of_float ?lx \ 0" by simp_all - have "(lb_cos prec (- ?lx)) \ cos (real (- ?lx))" + have "(lb_cos prec (- ?lx)) \ cos (real_of_float (- ?lx))" using lb_cos_minus[OF pi_lx lx_0] by simp also have "\ \ cos (x + (-k) * (2 * pi))" using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0] - by (simp only: uminus_float.rep_eq real_of_int_minus + by (simp only: uminus_float.rep_eq of_int_minus cos_minus mult_minus_left) simp finally have "(lb_cos prec (- ?lx)) \ cos x" unfolding cos_periodic_int . } @@ -1561,12 +1551,12 @@ { assume "0 \ ?lx" and pi_x: "x - k * (2 * pi) \ pi" with lx - have pi_lx: "?lx \ pi" and lx_0: "0 \ real ?lx" + have pi_lx: "?lx \ pi" and lx_0: "0 \ real_of_float ?lx" by auto have "cos (x + (-k) * (2 * pi)) \ cos ?lx" using cos_monotone_0_pi_le[OF lx_0 lx pi_x] - by (simp only: real_of_int_minus + by (simp only: of_int_minus cos_minus mult_minus_left) simp also have "\ \ (ub_cos prec ?lx)" using lb_cos[OF lx_0 pi_lx] by simp @@ -1576,12 +1566,12 @@ { assume pi_x: "- pi \ x - k * (2 * pi)" and "?ux \ 0" with ux - have pi_ux: "- pi \ ?ux" and ux_0: "real ?ux \ 0" + have pi_ux: "- pi \ ?ux" and ux_0: "real_of_float ?ux \ 0" by simp_all - have "cos (x + (-k) * (2 * pi)) \ cos (real (- ?ux))" + have "cos (x + (-k) * (2 * pi)) \ cos (real_of_float (- ?ux))" using cos_monotone_minus_pi_0'[OF pi_x ux ux_0] - by (simp only: uminus_float.rep_eq real_of_int_minus + by (simp only: uminus_float.rep_eq of_int_minus cos_minus mult_minus_left) simp also have "\ \ (ub_cos prec (- ?ux))" using lb_cos_minus[OF pi_ux ux_0, of prec] by simp @@ -1591,14 +1581,14 @@ { assume "?ux \ ?lpi" and x_ge_0: "0 \ x - k * (2 * pi)" with lpi ux - have pi_ux: "?ux \ pi" and ux_0: "0 \ real ?ux" + have pi_ux: "?ux \ pi" and ux_0: "0 \ real_of_float ?ux" by simp_all have "(lb_cos prec ?ux) \ cos ?ux" using lb_cos[OF ux_0 pi_ux] by simp also have "\ \ cos (x + (-k) * (2 * pi))" using cos_monotone_0_pi_le[OF x_ge_0 ux pi_ux] - by (simp only: real_of_int_minus + by (simp only: of_int_minus cos_minus mult_minus_left) simp finally have "(lb_cos prec ?ux) \ cos x" unfolding cos_periodic_int . } @@ -1648,7 +1638,7 @@ and u: "u = max (ub_cos prec ?lx) (ub_cos prec (- (?ux - 2 * ?lpi)))" by (auto simp add: bnds_cos_def Let_def) - have "cos x \ real u" + have "cos x \ real_of_float u" proof (cases "x - k * (2 * pi) < pi") case True hence "x - k * (2 * pi) \ pi" by simp @@ -1664,7 +1654,7 @@ hence "x - k * (2 * pi) - 2 * pi \ 0" using ux by simp - have ux_0: "real (?ux - 2 * ?lpi) \ 0" + have ux_0: "real_of_float (?ux - 2 * ?lpi) \ 0" using Cond by auto from 2 and Cond have "\ ?ux \ ?lpi" by auto @@ -1678,7 +1668,7 @@ unfolding cos_periodic_int .. also have "\ \ cos ((?ux - 2 * ?lpi))" using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0] - by (simp only: minus_float.rep_eq real_of_int_minus real_of_one + by (simp only: minus_float.rep_eq of_int_minus of_int_1 mult_minus_left mult_1_left) simp also have "\ = cos ((- (?ux - 2 * ?lpi)))" unfolding uminus_float.rep_eq cos_minus .. @@ -1711,7 +1701,7 @@ hence "0 \ x - k * (2 * pi) + 2 * pi" using lx by simp - have lx_0: "0 \ real (?lx + 2 * ?lpi)" + have lx_0: "0 \ real_of_float (?lx + 2 * ?lpi)" using Cond lpi by auto from 1 and Cond have "\ -?lpi \ ?lx" by auto @@ -1726,7 +1716,7 @@ unfolding cos_periodic_int .. also have "\ \ cos ((?lx + 2 * ?lpi))" using cos_monotone_0_pi_le[OF lx_0 lx_le_x pi_x] - by (simp only: minus_float.rep_eq real_of_int_minus real_of_one + by (simp only: minus_float.rep_eq of_int_minus of_int_1 mult_minus_left mult_1_left) simp also have "\ \ (ub_cos prec (?lx + 2 * ?lpi))" using lb_cos[OF lx_0 pi_lx] by simp @@ -1760,7 +1750,7 @@ (lapprox_rat prec 1 (int k)) (float_round_down prec (x * ub_exp_horner prec n (i + 1) (k * i) x))" lemma bnds_exp_horner: - assumes "real x \ 0" + assumes "real_of_float x \ 0" shows "exp x \ {lb_exp_horner prec (get_even n) 1 1 x .. ub_exp_horner prec (get_odd n) 1 1 x}" proof - have f_eq: "fact (Suc n) = fact n * ((\i::nat. i + 1) ^^ n) 1" for n @@ -1776,13 +1766,13 @@ have "lb_exp_horner prec (get_even n) 1 1 x \ exp x" proof - - have "lb_exp_horner prec (get_even n) 1 1 x \ (\j = 0.. (\j = 0.. \ exp x" proof - - obtain t where "\t\ \ \real x\" and "exp x = (\m = 0..t\ \ \real_of_float x\" and "exp x = (\m = 0.. exp t / (fact (get_even n)) * (real x) ^ (get_even n)" + moreover have "0 \ exp t / (fact (get_even n)) * (real_of_float x) ^ (get_even n)" by (auto simp: zero_le_even_power) ultimately show ?thesis using get_odd exp_gt_zero by auto qed @@ -1791,21 +1781,21 @@ moreover have "exp x \ ub_exp_horner prec (get_odd n) 1 1 x" proof - - have x_less_zero: "real x ^ get_odd n \ 0" - proof (cases "real x = 0") + have x_less_zero: "real_of_float x ^ get_odd n \ 0" + proof (cases "real_of_float x = 0") case True have "(get_odd n) \ 0" using get_odd[THEN odd_pos] by auto thus ?thesis unfolding True power_0_left by auto next - case False hence "real x < 0" using \real x \ 0\ by auto - show ?thesis by (rule less_imp_le, auto simp add: power_less_zero_eq \real x < 0\) + case False hence "real_of_float x < 0" using \real_of_float x \ 0\ by auto + show ?thesis by (rule less_imp_le, auto simp add: power_less_zero_eq \real_of_float x < 0\) qed - obtain t where "\t\ \ \real x\" - and "exp x = (\m = 0..t\ \ \real_of_float x\" + and "exp x = (\m = 0.. 0" + moreover have "exp t / (fact (get_odd n)) * (real_of_float x) ^ (get_odd n) \ 0" by (auto intro!: mult_nonneg_nonpos divide_nonpos_pos simp add: x_less_zero) - ultimately have "exp x \ (\j = 0.. (\j = 0.. \ ub_exp_horner prec (get_odd n) 1 1 x" using bounds(2) by auto @@ -1814,8 +1804,8 @@ ultimately show ?thesis by auto qed -lemma ub_exp_horner_nonneg: "real x \ 0 \ - 0 \ real (ub_exp_horner prec (get_odd n) (Suc 0) (Suc 0) x)" +lemma ub_exp_horner_nonneg: "real_of_float x \ 0 \ + 0 \ real_of_float (ub_exp_horner prec (get_odd n) (Suc 0) (Suc 0) x)" using bnds_exp_horner[of x prec n] by (intro order_trans[OF exp_ge_zero]) auto @@ -1850,7 +1840,7 @@ have "1 / 4 = (Float 1 (- 2))" unfolding Float_num by auto also have "\ \ lb_exp_horner 3 (get_even 3) 1 1 (- 1)" - by code_simp + by (subst less_eq_float.rep_eq [symmetric]) code_simp also have "\ \ exp (- 1 :: float)" using bnds_exp_horner[where x="- 1"] by auto finally show ?thesis @@ -1865,9 +1855,9 @@ let "?horner x" = "let y = ?lb_horner x in if y \ 0 then Float 1 (- 2) else y" have pos_horner: "0 < ?horner x" for x unfolding Let_def by (cases "?lb_horner x \ 0") auto - moreover have "0 < real ((?horner x) ^ num)" for x :: float and num :: nat + moreover have "0 < real_of_float ((?horner x) ^ num)" for x :: float and num :: nat proof - - have "0 < real (?horner x) ^ num" using \0 < ?horner x\ by simp + have "0 < real_of_float (?horner x) ^ num" using \0 < ?horner x\ by simp also have "\ = (?horner x) ^ num" by auto finally show ?thesis . qed @@ -1884,35 +1874,35 @@ let "?lb_exp_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" let "?ub_exp_horner x" = "ub_exp_horner prec (get_odd (prec + 2)) 1 1 x" - have "real x \ 0" and "\ x > 0" + have "real_of_float x \ 0" and "\ x > 0" using \x \ 0\ by auto show ?thesis proof (cases "x < - 1") case False - hence "- 1 \ real x" by auto + hence "- 1 \ real_of_float x" by auto show ?thesis proof (cases "?lb_exp_horner x \ 0") case True from \\ x < - 1\ - have "- 1 \ real x" by auto + have "- 1 \ real_of_float x" by auto hence "exp (- 1) \ exp x" unfolding exp_le_cancel_iff . from order_trans[OF exp_m1_ge_quarter this] have "Float 1 (- 2) \ exp x" unfolding Float_num . with True show ?thesis - using bnds_exp_horner \real x \ 0\ \\ x > 0\ \\ x < - 1\ by auto + using bnds_exp_horner \real_of_float x \ 0\ \\ x > 0\ \\ x < - 1\ by auto next case False thus ?thesis - using bnds_exp_horner \real x \ 0\ \\ x > 0\ \\ x < - 1\ by (auto simp add: Let_def) + using bnds_exp_horner \real_of_float x \ 0\ \\ x > 0\ \\ x < - 1\ by (auto simp add: Let_def) qed next case True let ?num = "nat (- int_floor_fl x)" - have "real (int_floor_fl x) < - 1" + have "real_of_int (int_floor_fl x) < - 1" using int_floor_fl[of x] \x < - 1\ by simp - hence "real (int_floor_fl x) < 0" by simp + hence "real_of_int (int_floor_fl x) < 0" by simp hence "int_floor_fl x < 0" by auto hence "1 \ - int_floor_fl x" by auto hence "0 < nat (- int_floor_fl x)" by auto @@ -1921,19 +1911,19 @@ have num_eq: "real ?num = - int_floor_fl x" using \0 < nat (- int_floor_fl x)\ by auto have "0 < - int_floor_fl x" - using \0 < ?num\[unfolded real_of_nat_less_iff[symmetric]] by simp - hence "real (int_floor_fl x) < 0" + using \0 < ?num\[unfolded of_nat_less_iff[symmetric]] by simp + hence "real_of_int (int_floor_fl x) < 0" unfolding less_float_def by auto - have fl_eq: "real (- int_floor_fl x) = real (- floor_fl x)" + have fl_eq: "real_of_int (- int_floor_fl x) = real_of_float (- floor_fl x)" by (simp add: floor_fl_def int_floor_fl_def) - from \0 < - int_floor_fl x\ have "0 \ real (- floor_fl x)" + from \0 < - int_floor_fl x\ have "0 \ real_of_float (- floor_fl x)" by (simp add: floor_fl_def int_floor_fl_def) - from \real (int_floor_fl x) < 0\ have "real (floor_fl x) < 0" + from \real_of_int (int_floor_fl x) < 0\ have "real_of_float (floor_fl x) < 0" by (simp add: floor_fl_def int_floor_fl_def) have "exp x \ ub_exp prec x" proof - - have div_less_zero: "real (float_divr prec x (- floor_fl x)) \ 0" - using float_divr_nonpos_pos_upper_bound[OF \real x \ 0\ \0 \ real (- floor_fl x)\] + have div_less_zero: "real_of_float (float_divr prec x (- floor_fl x)) \ 0" + using float_divr_nonpos_pos_upper_bound[OF \real_of_float x \ 0\ \0 \ real_of_float (- floor_fl x)\] unfolding less_eq_float_def zero_float.rep_eq . have "exp x = exp (?num * (x / ?num))" @@ -1946,7 +1936,7 @@ also have "\ \ (?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num" unfolding real_of_float_power by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto) - also have "\ \ real (power_up_fl prec (?ub_exp_horner (float_divr prec x (- floor_fl x))) ?num)" + also have "\ \ real_of_float (power_up_fl prec (?ub_exp_horner (float_divr prec x (- floor_fl x))) ?num)" by (auto simp add: real_power_up_fl intro!: power_up ub_exp_horner_nonneg div_less_zero) finally show ?thesis unfolding ub_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] floor_fl_def Let_def . @@ -1960,15 +1950,15 @@ show ?thesis proof (cases "?horner \ 0") case False - hence "0 \ real ?horner" by auto - - have div_less_zero: "real (float_divl prec x (- floor_fl x)) \ 0" - using \real (floor_fl x) < 0\ \real x \ 0\ + hence "0 \ real_of_float ?horner" by auto + + have div_less_zero: "real_of_float (float_divl prec x (- floor_fl x)) \ 0" + using \real_of_float (floor_fl x) < 0\ \real_of_float x \ 0\ by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg) have "(?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num \ exp (float_divl prec x (- floor_fl x)) ^ ?num" - using \0 \ real ?horner\[unfolded floor_fl_def[symmetric]] + using \0 \ real_of_float ?horner\[unfolded floor_fl_def[symmetric]] bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] by (auto intro!: power_mono) also have "\ \ exp (x / ?num) ^ ?num" @@ -1988,22 +1978,22 @@ have "power_down_fl prec (Float 1 (- 2)) ?num \ (Float 1 (- 2)) ^ ?num" by (metis Float_le_zero_iff less_imp_le linorder_not_less not_numeral_le_zero numeral_One power_down_fl) - then have "power_down_fl prec (Float 1 (- 2)) ?num \ real (Float 1 (- 2)) ^ ?num" + then have "power_down_fl prec (Float 1 (- 2)) ?num \ real_of_float (Float 1 (- 2)) ^ ?num" by simp also - have "real (floor_fl x) \ 0" and "real (floor_fl x) \ 0" - using \real (floor_fl x) < 0\ by auto - from divide_right_mono_neg[OF floor_fl[of x] \real (floor_fl x) \ 0\, unfolded divide_self[OF \real (floor_fl x) \ 0\]] + have "real_of_float (floor_fl x) \ 0" and "real_of_float (floor_fl x) \ 0" + using \real_of_float (floor_fl x) < 0\ by auto + from divide_right_mono_neg[OF floor_fl[of x] \real_of_float (floor_fl x) \ 0\, unfolded divide_self[OF \real_of_float (floor_fl x) \ 0\]] have "- 1 \ x / (- floor_fl x)" unfolding minus_float.rep_eq by auto from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]] have "Float 1 (- 2) \ exp (x / (- floor_fl x))" unfolding Float_num . - hence "real (Float 1 (- 2)) ^ ?num \ exp (x / (- floor_fl x)) ^ ?num" + hence "real_of_float (Float 1 (- 2)) ^ ?num \ exp (x / (- floor_fl x)) ^ ?num" by (metis Float_num(5) power_mono zero_le_divide_1_iff zero_le_numeral) also have "\ = exp x" unfolding num_eq fl_eq exp_real_of_nat_mult[symmetric] - using \real (floor_fl x) \ 0\ by auto + using \real_of_float (floor_fl x) \ 0\ by auto finally show ?thesis unfolding lb_exp.simps if_not_P[OF \\ 0 < x\] if_P[OF \x < - 1\] int_floor_fl_def Let_def if_P[OF True] real_of_float_power . @@ -2027,7 +2017,7 @@ have "lb_exp prec x \ exp x" proof - from exp_boundaries'[OF \-x \ 0\] - have ub_exp: "exp (- real x) \ ub_exp prec (-x)" + have ub_exp: "exp (- real_of_float x) \ ub_exp prec (-x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "float_divl prec 1 (ub_exp prec (-x)) \ 1 / ub_exp prec (-x)" @@ -2046,7 +2036,7 @@ have "\ 0 < -x" using \0 < x\ by auto from exp_boundaries'[OF \-x \ 0\] - have lb_exp: "lb_exp prec (-x) \ exp (- real x)" + have lb_exp: "lb_exp prec (-x) \ exp (- real_of_float x)" unfolding atLeastAtMost_iff minus_float.rep_eq by auto have "exp x \ (1 :: float) / lb_exp prec (-x)" @@ -2133,33 +2123,37 @@ qed lemma ln_float_bounds: - assumes "0 \ real x" - and "real x < 1" + assumes "0 \ real_of_float x" + and "real_of_float x < 1" shows "x * lb_ln_horner prec (get_even n) 1 x \ ln (x + 1)" (is "?lb \ ?ln") and "ln (x + 1) \ x * ub_ln_horner prec (get_odd n) 1 x" (is "?ln \ ?ub") proof - obtain ev where ev: "get_even n = 2 * ev" using get_even_double .. obtain od where od: "get_odd n = 2 * od + 1" using get_odd_double .. - let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real x)^(Suc n)" + let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real_of_float x)^(Suc n)" have "?lb \ setsum ?s {0 ..< 2 * ev}" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] - unfolding mult.commute[of "real x"] ev - using horner_bounds(1)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev", - OF \0 \ real x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real x\ + unfolding mult.commute[of "real_of_float x"] ev + using horner_bounds(1)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" + and lb="\n i k x. lb_ln_horner prec n k x" + and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev", + OF \0 \ real_of_float x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real_of_float x\ + unfolding real_of_float_power by (rule mult_right_mono) also have "\ \ ?ln" - using ln_bounds(1)[OF \0 \ real x\ \real x < 1\] by auto + using ln_bounds(1)[OF \0 \ real_of_float x\ \real_of_float x < 1\] by auto finally show "?lb \ ?ln" . have "?ln \ setsum ?s {0 ..< 2 * od + 1}" - using ln_bounds(2)[OF \0 \ real x\ \real x < 1\] by auto + using ln_bounds(2)[OF \0 \ real_of_float x\ \real_of_float x < 1\] by auto also have "\ \ ?ub" unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_left_distrib[symmetric] - unfolding mult.commute[of "real x"] od + unfolding mult.commute[of "real_of_float x"] od using horner_bounds(2)[where G="\ i k. Suc k" and F="\x. x" and f="\x. x" and lb="\n i k x. lb_ln_horner prec n k x" and ub="\n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1", - OF \0 \ real x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real x\ + OF \0 \ real_of_float x\ refl lb_ln_horner.simps ub_ln_horner.simps] \0 \ real_of_float x\ + unfolding real_of_float_power by (rule mult_right_mono) finally show "?ln \ ?ub" . qed @@ -2201,26 +2195,26 @@ have ln2_sum: "ln 2 = ln (1/2 + 1) + ln (1 / 3 + 1::real)" using ln_add[of "3 / 2" "1 / 2"] by auto have lb3: "?lthird \ 1 / 3" using lapprox_rat[of prec 1 3] by auto - hence lb3_ub: "real ?lthird < 1" by auto - have lb3_lb: "0 \ real ?lthird" using lapprox_rat_nonneg[of 1 3] by auto + hence lb3_ub: "real_of_float ?lthird < 1" by auto + have lb3_lb: "0 \ real_of_float ?lthird" using lapprox_rat_nonneg[of 1 3] by auto have ub3: "1 / 3 \ ?uthird" using rapprox_rat[of 1 3] by auto - hence ub3_lb: "0 \ real ?uthird" by auto - - have lb2: "0 \ real (Float 1 (- 1))" and ub2: "real (Float 1 (- 1)) < 1" + hence ub3_lb: "0 \ real_of_float ?uthird" by auto + + have lb2: "0 \ real_of_float (Float 1 (- 1))" and ub2: "real_of_float (Float 1 (- 1)) < 1" unfolding Float_num by auto have "0 \ (1::int)" and "0 < (3::int)" by auto - have ub3_ub: "real ?uthird < 1" + have ub3_ub: "real_of_float ?uthird < 1" by (simp add: Float.compute_rapprox_rat Float.compute_lapprox_rat rapprox_posrat_less1) have third_gt0: "(0 :: real) < 1 / 3 + 1" by auto - have uthird_gt0: "0 < real ?uthird + 1" using ub3_lb by auto - have lthird_gt0: "0 < real ?lthird + 1" using lb3_lb by auto + have uthird_gt0: "0 < real_of_float ?uthird + 1" using ub3_lb by auto + have lthird_gt0: "0 < real_of_float ?lthird + 1" using lb3_lb by auto show ?ub_ln2 unfolding ub_ln2_def Let_def ln2_sum Float_num(4)[symmetric] proof (rule float_plus_up_le, rule add_mono, fact ln_float_bounds(2)[OF lb2 ub2]) - have "ln (1 / 3 + 1) \ ln (real ?uthird + 1)" + have "ln (1 / 3 + 1) \ ln (real_of_float ?uthird + 1)" unfolding ln_le_cancel_iff[OF third_gt0 uthird_gt0] using ub3 by auto also have "\ \ ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird" using ln_float_bounds(2)[OF ub3_lb ub3_ub] . @@ -2230,7 +2224,7 @@ show ?lb_ln2 unfolding lb_ln2_def Let_def ln2_sum Float_num(4)[symmetric] proof (rule float_plus_down_le, rule add_mono, fact ln_float_bounds(1)[OF lb2 ub2]) - have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \ ln (real ?lthird + 1)" + have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \ ln (real_of_float ?lthird + 1)" using ln_float_bounds(1)[OF lb3_lb lb3_ub] . note float_round_down_le[OF this] also have "\ \ ln (1 / 3 + 1)" @@ -2265,18 +2259,18 @@ termination proof (relation "measure (\ v. let (prec, x) = case_sum id id v in (if x < 1 then 1 else 0))", auto) fix prec and x :: float - assume "\ real x \ 0" and "real x < 1" and "real (float_divl (max prec (Suc 0)) 1 x) < 1" - hence "0 < real x" "1 \ max prec (Suc 0)" "real x < 1" + assume "\ real_of_float x \ 0" and "real_of_float x < 1" and "real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1" + hence "0 < real_of_float x" "1 \ max prec (Suc 0)" "real_of_float x < 1" by auto - from float_divl_pos_less1_bound[OF \0 < real x\ \real x < 1\[THEN less_imp_le] \1 \ max prec (Suc 0)\] + from float_divl_pos_less1_bound[OF \0 < real_of_float x\ \real_of_float x < 1\[THEN less_imp_le] \1 \ max prec (Suc 0)\] show False - using \real (float_divl (max prec (Suc 0)) 1 x) < 1\ by auto + using \real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1\ by auto next fix prec x - assume "\ real x \ 0" and "real x < 1" and "real (float_divr prec 1 x) < 1" + assume "\ real_of_float x \ 0" and "real_of_float x < 1" and "real_of_float (float_divr prec 1 x) < 1" hence "0 < x" by auto - from float_divr_pos_less1_lower_bound[OF \0 < x\, of prec] \real x < 1\ show False - using \real (float_divr prec 1 x) < 1\ by auto + from float_divr_pos_less1_lower_bound[OF \0 < x\, of prec] \real_of_float x < 1\ show False + using \real_of_float (float_divr prec 1 x) < 1\ by auto qed lemma float_pos_eq_mantissa_pos: "x > 0 \ mantissa x > 0" @@ -2305,11 +2299,11 @@ unfolding zero_float_def[symmetric] using \0 < x\ by auto from denormalize_shift[OF assms(1) this] guess i . note i = this - have "2 powr (1 - (real (bitlen (mantissa x)) + real i)) = - 2 powr (1 - (real (bitlen (mantissa x)))) * inverse (2 powr (real i))" + have "2 powr (1 - (real_of_int (bitlen (mantissa x)) + real_of_int i)) = + 2 powr (1 - (real_of_int (bitlen (mantissa x)))) * inverse (2 powr (real i))" by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps) - hence "real (mantissa x) * 2 powr (1 - real (bitlen (mantissa x))) = - (real (mantissa x) * 2 ^ i) * 2 powr (1 - real (bitlen (mantissa x * 2 ^ i)))" + hence "real_of_int (mantissa x) * 2 powr (1 - real_of_int (bitlen (mantissa x))) = + (real_of_int (mantissa x) * 2 ^ i) * 2 powr (1 - real_of_int (bitlen (mantissa x * 2 ^ i)))" using \mantissa x > 0\ by (simp add: powr_realpow) then show ?th2 unfolding i by transfer auto @@ -2350,14 +2344,14 @@ proof - let ?B = "2^nat (bitlen m - 1)" def bl \ "bitlen m - 1" - have "0 < real m" and "\X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \ 0" + have "0 < real_of_int m" and "\X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \ 0" using assms by auto hence "0 \ bl" by (simp add: bitlen_def bl_def) show ?thesis proof (cases "0 \ e") case True thus ?thesis - unfolding bl_def[symmetric] using \0 < real m\ \0 \ bl\ + unfolding bl_def[symmetric] using \0 < real_of_int m\ \0 \ bl\ apply (simp add: ln_mult) apply (cases "e=0") apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr) @@ -2366,7 +2360,7 @@ next case False hence "0 < -e" by auto - have lne: "ln (2 powr real e) = ln (inverse (2 powr - e))" + have lne: "ln (2 powr real_of_int e) = ln (inverse (2 powr - e))" by (simp add: powr_minus) hence pow_gt0: "(0::real) < 2^nat (-e)" by auto @@ -2374,7 +2368,7 @@ by auto show ?thesis using False unfolding bl_def[symmetric] - using \0 < real m\ \0 \ bl\ + using \0 < real_of_int m\ \0 \ bl\ by (auto simp add: lne ln_mult ln_powr ln_div field_simps) qed qed @@ -2385,9 +2379,9 @@ (is "?lb \ ?ln \ ?ln \ ?ub") proof (cases "x < Float 1 1") case True - hence "real (x - 1) < 1" and "real x < 2" by auto + hence "real_of_float (x - 1) < 1" and "real_of_float x < 2" by auto have "\ x \ 0" and "\ x < 1" using \1 \ x\ by auto - hence "0 \ real (x - 1)" using \1 \ x\ by auto + hence "0 \ real_of_float (x - 1)" using \1 \ x\ by auto have [simp]: "(Float 3 (- 1)) = 3 / 2" by simp @@ -2397,7 +2391,7 @@ show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def - using ln_float_bounds[OF \0 \ real (x - 1)\ \real (x - 1) < 1\, of prec] + using ln_float_bounds[OF \0 \ real_of_float (x - 1)\ \real_of_float (x - 1) < 1\, of prec] \\ x \ 0\ \\ x < 1\ True by (auto intro!: float_round_down_le float_round_up_le) next @@ -2405,32 +2399,32 @@ hence *: "3 / 2 < x" by auto with ln_add[of "3 / 2" "x - 3 / 2"] - have add: "ln x = ln (3 / 2) + ln (real x * 2 / 3)" + have add: "ln x = ln (3 / 2) + ln (real_of_float x * 2 / 3)" by (auto simp add: algebra_simps diff_divide_distrib) let "?ub_horner x" = "float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x)" let "?lb_horner x" = "float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x)" - { have up: "real (rapprox_rat prec 2 3) \ 1" + { have up: "real_of_float (rapprox_rat prec 2 3) \ 1" by (rule rapprox_rat_le1) simp_all have low: "2 / 3 \ rapprox_rat prec 2 3" by (rule order_trans[OF _ rapprox_rat]) simp from mult_less_le_imp_less[OF * low] * - have pos: "0 < real (x * rapprox_rat prec 2 3 - 1)" by auto - - have "ln (real x * 2/3) - \ ln (real (x * rapprox_rat prec 2 3 - 1) + 1)" + have pos: "0 < real_of_float (x * rapprox_rat prec 2 3 - 1)" by auto + + have "ln (real_of_float x * 2/3) + \ ln (real_of_float (x * rapprox_rat prec 2 3 - 1) + 1)" proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) - show "real x * 2 / 3 \ real (x * rapprox_rat prec 2 3 - 1) + 1" + show "real_of_float x * 2 / 3 \ real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using * low by auto - show "0 < real x * 2 / 3" using * by simp - show "0 < real (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto + show "0 < real_of_float x * 2 / 3" using * by simp + show "0 < real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto qed also have "\ \ ?ub_horner (x * rapprox_rat prec 2 3 - 1)" proof (rule float_round_up_le, rule ln_float_bounds(2)) - from mult_less_le_imp_less[OF \real x < 2\ up] low * - show "real (x * rapprox_rat prec 2 3 - 1) < 1" by auto - show "0 \ real (x * rapprox_rat prec 2 3 - 1)" using pos by auto + from mult_less_le_imp_less[OF \real_of_float x < 2\ up] low * + show "real_of_float (x * rapprox_rat prec 2 3 - 1) < 1" by auto + show "0 \ real_of_float (x * rapprox_rat prec 2 3 - 1)" using pos by auto qed finally have "ln x \ ?ub_horner (Float 1 (-1)) + ?ub_horner ((x * rapprox_rat prec 2 3 - 1))" @@ -2444,23 +2438,23 @@ have up: "lapprox_rat prec 2 3 \ 2/3" by (rule order_trans[OF lapprox_rat], simp) - have low: "0 \ real (lapprox_rat prec 2 3)" + have low: "0 \ real_of_float (lapprox_rat prec 2 3)" using lapprox_rat_nonneg[of 2 3 prec] by simp have "?lb_horner ?max - \ ln (real ?max + 1)" + \ ln (real_of_float ?max + 1)" proof (rule float_round_down_le, rule ln_float_bounds(1)) - from mult_less_le_imp_less[OF \real x < 2\ up] * low - show "real ?max < 1" by (cases "real (lapprox_rat prec 2 3) = 0", + from mult_less_le_imp_less[OF \real_of_float x < 2\ up] * low + show "real_of_float ?max < 1" by (cases "real_of_float (lapprox_rat prec 2 3) = 0", auto simp add: real_of_float_max) - show "0 \ real ?max" by (auto simp add: real_of_float_max) + show "0 \ real_of_float ?max" by (auto simp add: real_of_float_max) qed - also have "\ \ ln (real x * 2/3)" + also have "\ \ ln (real_of_float x * 2/3)" proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) - show "0 < real ?max + 1" by (auto simp add: real_of_float_max) - show "0 < real x * 2/3" using * by auto - show "real ?max + 1 \ real x * 2/3" using * up - by (cases "0 < real x * real (lapprox_posrat prec 2 3) - 1", + show "0 < real_of_float ?max + 1" by (auto simp add: real_of_float_max) + show "0 < real_of_float x * 2/3" using * by auto + show "real_of_float ?max + 1 \ real_of_float x * 2/3" using * up + by (cases "0 < real_of_float x * real_of_float (lapprox_posrat prec 2 3) - 1", auto simp add: max_def) qed finally have "?lb_horner (Float 1 (- 1)) + ?lb_horner ?max \ ln x" @@ -2495,24 +2489,24 @@ have "1 \ Float m e" using \1 \ x\ Float unfolding less_eq_float_def by auto from bitlen_div[OF \0 < m\] float_gt1_scale[OF \1 \ Float m e\] \bl \ 0\ - have x_bnds: "0 \ real (?x - 1)" "real (?x - 1) < 1" + have x_bnds: "0 \ real_of_float (?x - 1)" "real_of_float (?x - 1) < 1" unfolding bl_def[symmetric] by (auto simp: powr_realpow[symmetric] field_simps inverse_eq_divide) (auto simp : powr_minus field_simps inverse_eq_divide) { have "float_round_down prec (lb_ln2 prec * ?s) \ ln 2 * (e + (bitlen m - 1))" - (is "real ?lb2 \ _") + (is "real_of_float ?lb2 \ _") apply (rule float_round_down_le) unfolding nat_0 power_0 mult_1_right times_float.rep_eq using lb_ln2[of prec] proof (rule mult_mono) from float_gt1_scale[OF \1 \ Float m e\] - show "0 \ real (Float (e + (bitlen m - 1)) 0)" by simp + show "0 \ real_of_float (Float (e + (bitlen m - 1)) 0)" by simp qed auto moreover from ln_float_bounds(1)[OF x_bnds] - have "float_round_down prec ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \ ln ?x" (is "real ?lb_horner \ _") + have "float_round_down prec ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \ ln ?x" (is "real_of_float ?lb_horner \ _") by (auto intro!: float_round_down_le) ultimately have "float_plus_down prec ?lb2 ?lb_horner \ ln x" unfolding Float ln_shifted_float[OF \0 < m\, of e] by (auto intro!: float_plus_down_le) @@ -2521,19 +2515,19 @@ { from ln_float_bounds(2)[OF x_bnds] have "ln ?x \ float_round_up prec ((?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1))" - (is "_ \ real ?ub_horner") + (is "_ \ real_of_float ?ub_horner") by (auto intro!: float_round_up_le) moreover have "ln 2 * (e + (bitlen m - 1)) \ float_round_up prec (ub_ln2 prec * ?s)" - (is "_ \ real ?ub2") + (is "_ \ real_of_float ?ub2") apply (rule float_round_up_le) unfolding nat_0 power_0 mult_1_right times_float.rep_eq using ub_ln2[of prec] proof (rule mult_mono) from float_gt1_scale[OF \1 \ Float m e\] - show "0 \ real (e + (bitlen m - 1))" by auto + show "0 \ real_of_int (e + (bitlen m - 1))" by auto have "0 \ ln (2 :: real)" by simp - thus "0 \ real (ub_ln2 prec)" using ub_ln2[of prec] by arith + thus "0 \ real_of_float (ub_ln2 prec)" using ub_ln2[of prec] by arith qed auto ultimately have "ln x \ float_plus_up prec ?ub2 ?ub_horner" unfolding Float ln_shifted_float[OF \0 < m\, of e] @@ -2562,29 +2556,29 @@ next case True have "\ x \ 0" using \0 < x\ by auto - from True have "real x \ 1" "x \ 1" + from True have "real_of_float x \ 1" "x \ 1" by simp_all - have "0 < real x" and "real x \ 0" + have "0 < real_of_float x" and "real_of_float x \ 0" using \0 < x\ by auto - hence A: "0 < 1 / real x" by auto + hence A: "0 < 1 / real_of_float x" by auto { let ?divl = "float_divl (max prec 1) 1 x" - have A': "1 \ ?divl" using float_divl_pos_less1_bound[OF \0 < real x\ \real x \ 1\] by auto - hence B: "0 < real ?divl" by auto + have A': "1 \ ?divl" using float_divl_pos_less1_bound[OF \0 < real_of_float x\ \real_of_float x \ 1\] by auto + hence B: "0 < real_of_float ?divl" by auto have "ln ?divl \ ln (1 / x)" unfolding ln_le_cancel_iff[OF B A] using float_divl[of _ 1 x] by auto - hence "ln x \ - ln ?divl" unfolding nonzero_inverse_eq_divide[OF \real x \ 0\, symmetric] ln_inverse[OF \0 < real x\] by auto + hence "ln x \ - ln ?divl" unfolding nonzero_inverse_eq_divide[OF \real_of_float x \ 0\, symmetric] ln_inverse[OF \0 < real_of_float x\] by auto from this ub_ln_lb_ln_bounds'[OF A', THEN conjunct1, THEN le_imp_neg_le] have "?ln \ - the (lb_ln prec ?divl)" unfolding uminus_float.rep_eq by (rule order_trans) } moreover { let ?divr = "float_divr prec 1 x" have A': "1 \ ?divr" using float_divr_pos_less1_lower_bound[OF \0 < x\ \x \ 1\] unfolding less_eq_float_def less_float_def by auto - hence B: "0 < real ?divr" by auto + hence B: "0 < real_of_float ?divr" by auto have "ln (1 / x) \ ln ?divr" unfolding ln_le_cancel_iff[OF A B] using float_divr[of 1 x] by auto - hence "- ln ?divr \ ln x" unfolding nonzero_inverse_eq_divide[OF \real x \ 0\, symmetric] ln_inverse[OF \0 < real x\] by auto + hence "- ln ?divr \ ln x" unfolding nonzero_inverse_eq_divide[OF \real_of_float x \ 0\, symmetric] ln_inverse[OF \0 < real_of_float x\] by auto from ub_ln_lb_ln_bounds'[OF A', THEN conjunct2, THEN le_imp_neg_le] this have "- the (ub_ln prec ?divr) \ ?ln" unfolding uminus_float.rep_eq by (rule order_trans) } @@ -2594,7 +2588,7 @@ lemma lb_ln: assumes "Some y = lb_ln prec x" - shows "y \ ln x" and "0 < real x" + shows "y \ ln x" and "0 < real_of_float x" proof - have "0 < x" proof (rule ccontr) @@ -2604,7 +2598,7 @@ thus False using assms by auto qed - thus "0 < real x" by auto + thus "0 < real_of_float x" by auto have "the (lb_ln prec x) \ ln x" using ub_ln_lb_ln_bounds[OF \0 < x\] .. thus "y \ ln x" @@ -2613,7 +2607,7 @@ lemma ub_ln: assumes "Some y = ub_ln prec x" - shows "ln x \ y" and "0 < real x" + shows "ln x \ y" and "0 < real_of_float x" proof - have "0 < x" proof (rule ccontr) @@ -2622,7 +2616,7 @@ thus False using assms by auto qed - thus "0 < real x" by auto + thus "0 < real_of_float x" by auto have "ln x \ the (ub_ln prec x)" using ub_ln_lb_ln_bounds[OF \0 < x\] .. thus "ln x \ y" @@ -2638,16 +2632,16 @@ hence l: "Some l = lb_ln prec lx " and u: "Some u = ub_ln prec ux" and x: "x \ {lx .. ux}" by auto - have "ln ux \ u" and "0 < real ux" + have "ln ux \ u" and "0 < real_of_float ux" using ub_ln u by auto - have "l \ ln lx" and "0 < real lx" and "0 < x" + have "l \ ln lx" and "0 < real_of_float lx" and "0 < x" using lb_ln[OF l] x by auto - from ln_le_cancel_iff[OF \0 < real lx\ \0 < x\] \l \ ln lx\ + from ln_le_cancel_iff[OF \0 < real_of_float lx\ \0 < x\] \l \ ln lx\ have "l \ ln x" using x unfolding atLeastAtMost_iff by auto moreover - from ln_le_cancel_iff[OF \0 < x\ \0 < real ux\] \ln ux \ real u\ + from ln_le_cancel_iff[OF \0 < x\ \0 < real_of_float ux\] \ln ux \ real_of_float u\ have "ln x \ u" using x unfolding atLeastAtMost_iff by auto ultimately show "l \ ln x \ ln x \ u" .. @@ -2746,19 +2740,20 @@ "lift_un' (Some (l1, u1)) f = Some (f l1 u1)" | "lift_un' b f = None" -definition "bounded_by xs vs \ +definition bounded_by :: "real list \ (float \ float) option list \ bool" where + "bounded_by xs vs \ (\ i < length vs. case vs ! i of None \ True - | Some (l, u) \ xs ! i \ { real l .. real u })" - + | Some (l, u) \ xs ! i \ { real_of_float l .. real_of_float u })" + lemma bounded_byE: assumes "bounded_by xs vs" shows "\ i. i < length vs \ case vs ! i of None \ True - | Some (l, u) \ xs ! i \ { real l .. real u }" + | Some (l, u) \ xs ! i \ { real_of_float l .. real_of_float u }" using assms bounded_by_def by blast lemma bounded_by_update: assumes "bounded_by xs vs" - and bnd: "xs ! i \ { real l .. real u }" + and bnd: "xs ! i \ { real_of_float l .. real_of_float u }" shows "bounded_by xs (vs[i := Some (l,u)])" proof - { @@ -2766,7 +2761,7 @@ let ?vs = "vs[i := Some (l,u)]" assume "j < length ?vs" hence [simp]: "j < length vs" by simp - have "case ?vs ! j of None \ True | Some (l, u) \ xs ! j \ { real l .. real u }" + have "case ?vs ! j of None \ True | Some (l, u) \ xs ! j \ { real_of_float l .. real_of_float u }" proof (cases "?vs ! j") case (Some b) thus ?thesis @@ -2949,7 +2944,7 @@ and lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f" and Pa: "\l u. Some (l, u) = approx prec a bs \ l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - shows "real l \ f' (interpret_floatarith a xs) \ f' (interpret_floatarith a xs) \ real u" + shows "real_of_float l \ f' (interpret_floatarith a xs) \ f' (interpret_floatarith a xs) \ real_of_float u" proof - from lift_un'[OF lift_un'_Some Pa] obtain l1 u1 where "l1 \ interpret_floatarith a xs" @@ -3039,7 +3034,7 @@ and lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f" and Pa: "\l u. Some (l, u) = approx prec a bs \ l \ interpret_floatarith a xs \ interpret_floatarith a xs \ u" - shows "real l \ f' (interpret_floatarith a xs) \ f' (interpret_floatarith a xs) \ real u" + shows "real_of_float l \ f' (interpret_floatarith a xs) \ f' (interpret_floatarith a xs) \ real_of_float u" proof - from lift_un[OF lift_un_Some Pa] obtain l1 u1 where "l1 \ interpret_floatarith a xs" @@ -3109,37 +3104,37 @@ show False using l' unfolding if_not_P[OF P] by auto qed - moreover have l1_le_u1: "real l1 \ real u1" + moreover have l1_le_u1: "real_of_float l1 \ real_of_float u1" using l1 u1 by auto - ultimately have "real l1 \ 0" and "real u1 \ 0" + ultimately have "real_of_float l1 \ 0" and "real_of_float u1 \ 0" by auto have inv: "inverse u1 \ inverse (interpret_floatarith a xs) \ inverse (interpret_floatarith a xs) \ inverse l1" proof (cases "0 < l1") case True - hence "0 < real u1" and "0 < real l1" "0 < interpret_floatarith a xs" + hence "0 < real_of_float u1" and "0 < real_of_float l1" "0 < interpret_floatarith a xs" using l1_le_u1 l1 by auto show ?thesis - unfolding inverse_le_iff_le[OF \0 < real u1\ \0 < interpret_floatarith a xs\] - inverse_le_iff_le[OF \0 < interpret_floatarith a xs\ \0 < real l1\] + unfolding inverse_le_iff_le[OF \0 < real_of_float u1\ \0 < interpret_floatarith a xs\] + inverse_le_iff_le[OF \0 < interpret_floatarith a xs\ \0 < real_of_float l1\] using l1 u1 by auto next case False hence "u1 < 0" using either by blast - hence "real u1 < 0" and "real l1 < 0" "interpret_floatarith a xs < 0" + hence "real_of_float u1 < 0" and "real_of_float l1 < 0" "interpret_floatarith a xs < 0" using l1_le_u1 u1 by auto show ?thesis - unfolding inverse_le_iff_le_neg[OF \real u1 < 0\ \interpret_floatarith a xs < 0\] - inverse_le_iff_le_neg[OF \interpret_floatarith a xs < 0\ \real l1 < 0\] + unfolding inverse_le_iff_le_neg[OF \real_of_float u1 < 0\ \interpret_floatarith a xs < 0\] + inverse_le_iff_le_neg[OF \interpret_floatarith a xs < 0\ \real_of_float l1 < 0\] using l1 u1 by auto qed from l' have "l = float_divl prec 1 u1" by (cases "0 < l1 \ u1 < 0") auto hence "l \ inverse u1" - unfolding nonzero_inverse_eq_divide[OF \real u1 \ 0\] + unfolding nonzero_inverse_eq_divide[OF \real_of_float u1 \ 0\] using float_divl[of prec 1 u1] by auto also have "\ \ inverse (interpret_floatarith a xs)" using inv by auto @@ -3148,7 +3143,7 @@ from u' have "u = float_divr prec 1 l1" by (cases "0 < l1 \ u1 < 0") auto hence "inverse l1 \ u" - unfolding nonzero_inverse_eq_divide[OF \real l1 \ 0\] + unfolding nonzero_inverse_eq_divide[OF \real_of_float l1 \ 0\] using float_divr[of 1 l1 prec] by auto hence "inverse (interpret_floatarith a xs) \ u" by (rule order_trans[OF inv[THEN conjunct2]]) @@ -3274,7 +3269,7 @@ case (Suc s) let ?m = "(l + u) * Float 1 (- 1)" - have "real l \ ?m" and "?m \ real u" + have "real_of_float l \ ?m" and "?m \ real_of_float u" unfolding less_eq_float_def using Suc.prems by auto with \x \ { l .. u }\ @@ -3355,7 +3350,7 @@ then obtain l u l' u' where l_eq: "Some (l, u) = approx prec a vs" and u_eq: "Some (l', u') = approx prec b vs" - and inequality: "real (float_plus_up prec u (-l')) < 0" + and inequality: "real_of_float (float_plus_up prec u (-l')) < 0" by (cases "approx prec a vs", auto, cases "approx prec b vs", auto) from le_less_trans[OF float_plus_up inequality] approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq] @@ -3365,7 +3360,7 @@ then obtain l u l' u' where l_eq: "Some (l, u) = approx prec a vs" and u_eq: "Some (l', u') = approx prec b vs" - and inequality: "real (float_plus_up prec u (-l')) \ 0" + and inequality: "real_of_float (float_plus_up prec u (-l')) \ 0" by (cases "approx prec a vs", auto, cases "approx prec b vs", auto) from order_trans[OF float_plus_up inequality] approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq] @@ -3376,7 +3371,7 @@ where x_eq: "Some (lx, ux) = approx prec x vs" and l_eq: "Some (l, u) = approx prec a vs" and u_eq: "Some (l', u') = approx prec b vs" - and inequality: "real (float_plus_up prec u (-lx)) \ 0" "real (float_plus_up prec ux (-l')) \ 0" + and inequality: "real_of_float (float_plus_up prec u (-lx)) \ 0" "real_of_float (float_plus_up prec ux (-l')) \ 0" by (cases "approx prec x vs", auto, cases "approx prec a vs", auto, cases "approx prec b vs", auto) @@ -3452,7 +3447,7 @@ next case (Power a n) thus ?case - by (cases n) (auto intro!: derivative_eq_intros simp del: power_Suc simp add: real_of_nat_def) + by (cases n) (auto intro!: derivative_eq_intros simp del: power_Suc) next case (Ln a) thus ?case by (auto intro!: derivative_eq_intros simp add: divide_inverse) @@ -3522,7 +3517,7 @@ lemma bounded_by_update_var: assumes "bounded_by xs vs" and "vs ! i = Some (l, u)" - and bnd: "x \ { real l .. real u }" + and bnd: "x \ { real_of_float l .. real_of_float u }" shows "bounded_by (xs[i := x]) vs" proof (cases "i < length xs") case False @@ -3532,7 +3527,7 @@ case True let ?xs = "xs[i := x]" from True have "i < length ?xs" by auto - have "case vs ! j of None \ True | Some (l, u) \ ?xs ! j \ {real l .. real u}" + have "case vs ! j of None \ True | Some (l, u) \ ?xs ! j \ {real_of_float l .. real_of_float u}" if "j < length vs" for j proof (cases "vs ! j") case None @@ -3557,7 +3552,7 @@ lemma isDERIV_approx': assumes "bounded_by xs vs" and vs_x: "vs ! x = Some (l, u)" - and X_in: "X \ {real l .. real u}" + and X_in: "X \ {real_of_float l .. real_of_float u}" and approx: "isDERIV_approx prec x f vs" shows "isDERIV x f (xs[x := X])" proof - @@ -3612,10 +3607,10 @@ lemma bounded_by_Cons: assumes bnd: "bounded_by xs vs" - and x: "x \ { real l .. real u }" + and x: "x \ { real_of_float l .. real_of_float u }" shows "bounded_by (x#xs) ((Some (l, u))#vs)" proof - - have "case ((Some (l,u))#vs) ! i of Some (l, u) \ (x#xs)!i \ { real l .. real u } | None \ True" + have "case ((Some (l,u))#vs) ! i of Some (l, u) \ (x#xs)!i \ { real_of_float l .. real_of_float u } | None \ True" if *: "i < length ((Some (l, u))#vs)" for i proof (cases i) case 0 @@ -3689,7 +3684,7 @@ from approx[OF this a] have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \ { l1 .. u1 }" - (is "?f 0 (real c) \ _") + (is "?f 0 (real_of_float c) \ _") by auto have funpow_Suc[symmetric]: "(f ^^ Suc n) x = (f ^^ n) (f x)" @@ -3698,7 +3693,7 @@ from Suc.hyps[OF ate, unfolded this] obtain n where DERIV_hyp: "\m z. \ m < n ; (z::real) \ { lx .. ux } \ \ DERIV (?f (Suc m)) z :> ?f (Suc (Suc m)) z" - and hyp: "\t \ {real lx .. real ux}. + and hyp: "\t \ {real_of_float lx .. real_of_float ux}. (\ i = 0.. j \ {Suc k.. j \ {Suc k.. {l2 .. u2}" (is "\ t \ _. ?X (Suc k) f n t \ _") @@ -3737,9 +3732,9 @@ have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some (l1, u1), Some (l2, u2), vs!x]" by (auto intro!: bounded_by_Cons) from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]] - have "?X (Suc k) f n t * (xs!x - real c) * inverse k + ?f 0 c \ {l .. u}" + have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \ {l .. u}" by (auto simp add: algebra_simps) - also have "?X (Suc k) f n t * (xs!x - real c) * inverse (real k) + ?f 0 c = + also have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse (real k) + ?f 0 c = (\ i = 0.. j \ {k.. j \ {k.. {1..<1 + k}) = fact (k :: nat)" - using fact_altdef_nat Suc_eq_plus1_left atLeastLessThanSuc_atLeastAtMost real_fact_nat - by presburger +by (metis Suc_eq_plus1_left atLeastLessThanSuc_atLeastAtMost fact_altdef_nat of_nat_fact) lemma approx_tse: assumes "bounded_by xs vs" and bnd_x: "vs ! x = Some (lx, ux)" - and bnd_c: "real c \ {lx .. ux}" + and bnd_c: "real_of_float c \ {lx .. ux}" and "x < length vs" and "x < length xs" and ate: "Some (l, u) = approx_tse prec x s c 1 f vs" shows "interpret_floatarith f xs \ {l .. u}" @@ -3772,7 +3766,7 @@ from approx_tse_generic[OF \bounded_by xs vs\ this bnd_x ate] obtain n - where DERIV: "\ m z. m < n \ real lx \ z \ z \ real ux \ DERIV (F m) z :> F (Suc m) z" + where DERIV: "\ m z. m < n \ real_of_float lx \ z \ z \ real_of_float ux \ DERIV (F m) z :> F (Suc m) z" and hyp: "\ (t::real). t \ {lx .. ux} \ (\ j = 0.. real c" "real c \ ux" "lx \ xs!x" "xs!x \ ux" + have "lx \ real_of_float c" "real_of_float c \ ux" "lx \ xs!x" "xs!x \ ux" using Suc bnd_c \bounded_by xs vs\[THEN bounded_byE, OF \x < length vs\] bnd_x by auto from Taylor.taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False] obtain t::real where t_bnd: "if xs ! x < c then xs ! x < t \ t < c else c < t \ t < xs ! x" @@ -3833,7 +3827,7 @@ fixes x :: real assumes "approx_tse_form' prec t f s l u cmp" and "x \ {l .. u}" - shows "\l' u' ly uy. x \ {l' .. u'} \ real l \ l' \ u' \ real u \ cmp ly uy \ + shows "\l' u' ly uy. x \ {l' .. u'} \ real_of_float l \ l' \ u' \ real_of_float u \ cmp ly uy \ approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)" using assms proof (induct s arbitrary: l u) @@ -3850,7 +3844,7 @@ and u: "approx_tse_form' prec t f s ?m u cmp" by (auto simp add: Let_def lazy_conj) - have m_l: "real l \ ?m" and m_u: "?m \ real u" + have m_l: "real_of_float l \ ?m" and m_u: "?m \ real_of_float u" unfolding less_eq_float_def using Suc.prems by auto with \x \ { l .. u }\ consider "x \ { l .. ?m}" | "x \ {?m .. u}" by atomize_elim auto @@ -3859,7 +3853,7 @@ case 1 from Suc.hyps[OF l this] obtain l' u' ly uy where - "x \ {l' .. u'} \ real l \ l' \ real u' \ ?m \ cmp ly uy \ + "x \ {l' .. u'} \ real_of_float l \ l' \ real_of_float u' \ ?m \ cmp ly uy \ approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)" by blast with m_u show ?thesis @@ -3868,7 +3862,7 @@ case 2 from Suc.hyps[OF u this] obtain l' u' ly uy where - "x \ { l' .. u' } \ ?m \ real l' \ u' \ real u \ cmp ly uy \ + "x \ { l' .. u' } \ ?m \ real_of_float l' \ u' \ real_of_float u \ cmp ly uy \ approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)" by blast with m_u show ?thesis @@ -3885,8 +3879,8 @@ from approx_tse_form'[OF tse x] obtain l' u' ly uy where x': "x \ {l' .. u'}" - and "l \ real l'" - and "real u' \ u" and "0 < ly" + and "real_of_float l \ real_of_float l'" + and "real_of_float u' \ real_of_float u" and "0 < ly" and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)" by blast @@ -3908,8 +3902,8 @@ from approx_tse_form'[OF tse x] obtain l' u' ly uy where x': "x \ {l' .. u'}" - and "l \ real l'" - and "real u' \ u" and "0 \ ly" + and "l \ real_of_float l'" + and "real_of_float u' \ u" and "0 \ ly" and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)" by blast diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Decision_Procs/Ferrack.thy --- a/src/HOL/Decision_Procs/Ferrack.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Decision_Procs/Ferrack.thy Tue Nov 10 14:18:41 2015 +0000 @@ -30,13 +30,13 @@ (* Semantics of numeral terms (num) *) primrec Inum :: "real list \ num \ real" where - "Inum bs (C c) = (real c)" + "Inum bs (C c) = (real_of_int c)" | "Inum bs (Bound n) = bs!n" -| "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" +| "Inum bs (CN n c a) = (real_of_int c) * (bs!n) + (Inum bs a)" | "Inum bs (Neg a) = -(Inum bs a)" | "Inum bs (Add a b) = Inum bs a + Inum bs b" | "Inum bs (Sub a b) = Inum bs a - Inum bs b" -| "Inum bs (Mul c a) = (real c) * Inum bs a" +| "Inum bs (Mul c a) = (real_of_int c) * Inum bs a" (* FORMULAE *) datatype fm = T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| @@ -518,7 +518,7 @@ lemma reducecoeffh: assumes gt: "dvdnumcoeff t g" and gp: "g > 0" - shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" + shows "real_of_int g *(Inum bs (reducecoeffh t g)) = Inum bs t" using gt proof (induct t rule: reducecoeffh.induct) case (1 i) @@ -618,7 +618,7 @@ from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . qed -lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" +lemma reducecoeff: "real_of_int (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" proof - let ?g = "numgcd t" have "?g \ 0" @@ -778,8 +778,8 @@ else (t', n))))" lemma simp_num_pair_ci: - shows "((\(t,n). Inum bs t / real n) (simp_num_pair (t,n))) = - ((\(t,n). Inum bs t / real n) (t, n))" + shows "((\(t,n). Inum bs t / real_of_int n) (simp_num_pair (t,n))) = + ((\(t,n). Inum bs t / real_of_int n) (t, n))" (is "?lhs = ?rhs") proof - let ?t' = "simpnum t" @@ -819,15 +819,15 @@ have gpdd: "?g' dvd n" by simp have gpdgp: "?g' dvd ?g'" by simp from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] - have th2:"real ?g' * ?t = Inum bs ?t'" + have th2:"real_of_int ?g' * ?t = Inum bs ?t'" by simp - from g1 g'1 have "?lhs = ?t / real (n div ?g')" + from g1 g'1 have "?lhs = ?t / real_of_int (n div ?g')" by (simp add: simp_num_pair_def Let_def) - also have "\ = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" + also have "\ = (real_of_int ?g' * ?t) / (real_of_int ?g' * (real_of_int (n div ?g')))" by simp - also have "\ = (Inum bs ?t' / real n)" + also have "\ = (Inum bs ?t' / real_of_int n)" using real_of_int_div[OF gpdd] th2 gp0 by simp - finally have "?lhs = Inum bs t / real n" + finally have "?lhs = Inum bs t / real_of_int n" by simp then show ?thesis by (simp add: simp_num_pair_def) @@ -1278,17 +1278,17 @@ next case (3 c e) from 3 have nb: "numbound0 e" by simp - from 3 have cp: "real c > 0" by simp + from 3 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x < ?z" - then have "(real c * x < - ?e)" + then have "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - then have "real c * x + ?e < 0" by arith - then have "real c * x + ?e \ 0" by simp + then have "real_of_int c * x + ?e < 0" by arith + then have "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1297,17 +1297,17 @@ next case (4 c e) from 4 have nb: "numbound0 e" by simp - from 4 have cp: "real c > 0" by simp + from 4 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x < ?z" - then have "(real c * x < - ?e)" + then have "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - then have "real c * x + ?e < 0" by arith - then have "real c * x + ?e \ 0" by simp + then have "real_of_int c * x + ?e < 0" by arith + then have "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1316,16 +1316,16 @@ next case (5 c e) from 5 have nb: "numbound0 e" by simp - from 5 have cp: "real c > 0" by simp + from 5 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x < ?z" - then have "(real c * x < - ?e)" + then have "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - then have "real c * x + ?e < 0" by arith + then have "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1334,16 +1334,16 @@ next case (6 c e) from 6 have nb: "numbound0 e" by simp - from lp 6 have cp: "real c > 0" by simp + from lp 6 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x < ?z" - then have "(real c * x < - ?e)" + then have "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - then have "real c * x + ?e < 0" by arith + then have "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1352,16 +1352,16 @@ next case (7 c e) from 7 have nb: "numbound0 e" by simp - from 7 have cp: "real c > 0" by simp + from 7 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x < ?z" - then have "(real c * x < - ?e)" + then have "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - then have "real c * x + ?e < 0" by arith + then have "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1370,16 +1370,16 @@ next case (8 c e) from 8 have nb: "numbound0 e" by simp - from 8 have cp: "real c > 0" by simp + from 8 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x < ?z" - then have "(real c * x < - ?e)" + then have "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - then have "real c * x + ?e < 0" by arith + then have "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1408,17 +1408,17 @@ next case (3 c e) from 3 have nb: "numbound0 e" by simp - from 3 have cp: "real c > 0" by simp + from 3 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - then have "real c * x + ?e > 0" by arith - then have "real c * x + ?e \ 0" by simp + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + then have "real_of_int c * x + ?e > 0" by arith + then have "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1427,17 +1427,17 @@ next case (4 c e) from 4 have nb: "numbound0 e" by simp - from 4 have cp: "real c > 0" by simp + from 4 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - then have "real c * x + ?e > 0" by arith - then have "real c * x + ?e \ 0" by simp + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + then have "real_of_int c * x + ?e > 0" by arith + then have "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1446,16 +1446,16 @@ next case (5 c e) from 5 have nb: "numbound0 e" by simp - from 5 have cp: "real c > 0" by simp + from 5 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - then have "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + then have "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1464,16 +1464,16 @@ next case (6 c e) from 6 have nb: "numbound0 e" by simp - from 6 have cp: "real c > 0" by simp + from 6 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - then have "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + then have "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1482,16 +1482,16 @@ next case (7 c e) from 7 have nb: "numbound0 e" by simp - from 7 have cp: "real c > 0" by simp + from 7 have cp: "real_of_int c > 0" by simp fix a let ?e = "Inum (a # bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - then have "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + then have "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1500,16 +1500,16 @@ next case (8 c e) from 8 have nb: "numbound0 e" by simp - from 8 have cp: "real c > 0" by simp + from 8 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" { fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - then have "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + then have "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -1581,10 +1581,10 @@ lemma usubst_I: assumes lp: "isrlfm p" - and np: "real n > 0" + and np: "real_of_int n > 0" and nbt: "numbound0 t" shows "(Ifm (x # bs) (usubst p (t,n)) = - Ifm (((Inum (x # bs) t) / (real n)) # bs) p) \ bound0 (usubst p (t, n))" + Ifm (((Inum (x # bs) t) / (real_of_int n)) # bs) p) \ bound0 (usubst p (t, n))" (is "(?I x (usubst p (t, n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") @@ -1592,65 +1592,65 @@ proof (induct p rule: usubst.induct) case (5 c e) with assms have cp: "c > 0" and nb: "numbound0 e" by simp_all - have "?I ?u (Lt (CN 0 c e)) \ real c * (?t / ?n) + ?N x e < 0" + have "?I ?u (Lt (CN 0 c e)) \ real_of_int c * (?t / ?n) + ?N x e < 0" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ \ ?n * (real c * (?t / ?n)) + ?n*(?N x e) < 0" - by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ \ ?n * (real_of_int c * (?t / ?n)) + ?n*(?N x e) < 0" + by (simp only: pos_less_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ \ real c * ?t + ?n * (?N x e) < 0" using np by simp + also have "\ \ real_of_int c * ?t + ?n * (?N x e) < 0" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (6 c e) with assms have cp: "c > 0" and nb: "numbound0 e" by simp_all - have "?I ?u (Le (CN 0 c e)) \ real c * (?t / ?n) + ?N x e \ 0" + have "?I ?u (Le (CN 0 c e)) \ real_of_int c * (?t / ?n) + ?N x e \ 0" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_le_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" using np by simp + also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (7 c e) with assms have cp: "c >0" and nb: "numbound0 e" by simp_all - have "?I ?u (Gt (CN 0 c e)) \ real c *(?t / ?n) + ?N x e > 0" + have "?I ?u (Gt (CN 0 c e)) \ real_of_int c *(?t / ?n) + ?N x e > 0" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ \ ?n * (real c * (?t / ?n)) + ?n * ?N x e > 0" - by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ \ ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x e > 0" + by (simp only: pos_divide_less_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ \ real c * ?t + ?n * ?N x e > 0" using np by simp + also have "\ \ real_of_int c * ?t + ?n * ?N x e > 0" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (8 c e) with assms have cp: "c > 0" and nb: "numbound0 e" by simp_all - have "?I ?u (Ge (CN 0 c e)) \ real c * (?t / ?n) + ?N x e \ 0" + have "?I ?u (Ge (CN 0 c e)) \ real_of_int c * (?t / ?n) + ?N x e \ 0" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ \ ?n * (real c * (?t / ?n)) + ?n * ?N x e \ 0" - by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ \ ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x e \ 0" + by (simp only: pos_divide_le_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ \ real c * ?t + ?n * ?N x e \ 0" using np by simp + also have "\ \ real_of_int c * ?t + ?n * ?N x e \ 0" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (3 c e) with assms have cp: "c > 0" and nb: "numbound0 e" by simp_all - from np have np: "real n \ 0" by simp - have "?I ?u (Eq (CN 0 c e)) \ real c * (?t / ?n) + ?N x e = 0" + from np have np: "real_of_int n \ 0" by simp + have "?I ?u (Eq (CN 0 c e)) \ real_of_int c * (?t / ?n) + ?N x e = 0" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ \ ?n * (real c * (?t / ?n)) + ?n * ?N x e = 0" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ \ ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x e = 0" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ \ real c * ?t + ?n * ?N x e = 0" using np by simp + also have "\ \ real_of_int c * ?t + ?n * ?N x e = 0" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (4 c e) with assms have cp: "c >0" and nb: "numbound0 e" by simp_all - from np have np: "real n \ 0" by simp - have "?I ?u (NEq (CN 0 c e)) \ real c * (?t / ?n) + ?N x e \ 0" + from np have np: "real_of_int n \ 0" by simp + have "?I ?u (NEq (CN 0 c e)) \ real_of_int c * (?t / ?n) + ?N x e \ 0" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ \ ?n * (real c * (?t / ?n)) + ?n * ?N x e \ 0" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ \ ?n * (real_of_int c * (?t / ?n)) + ?n * ?N x e \ 0" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ \ real c * ?t + ?n * ?N x e \ 0" using np by simp + also have "\ \ real_of_int c * ?t + ?n * ?N x e \ 0" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) -qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"]) +qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real_of_int n" and b'="x"]) lemma uset_l: assumes lp: "isrlfm p" @@ -1661,18 +1661,18 @@ assumes lp: "isrlfm p" and nmi: "\ (Ifm (a # bs) (minusinf p))" (is "\ (Ifm (a # bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\(s,m) \ set (uset p). x \ Inum (a#bs) s / real m" - (is "\(s,m) \ ?U p. x \ ?N a s / real m") + shows "\(s,m) \ set (uset p). x \ Inum (a#bs) s / real_of_int m" + (is "\(s,m) \ ?U p. x \ ?N a s / real_of_int m") proof - - have "\(s,m) \ set (uset p). real m * x \ Inum (a#bs) s" - (is "\(s,m) \ ?U p. real m *x \ ?N a s") + have "\(s,m) \ set (uset p). real_of_int m * x \ Inum (a#bs) s" + (is "\(s,m) \ ?U p. real_of_int m *x \ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct) (auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) - then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" + then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real_of_int m * x \ ?N a s" by blast - from uset_l[OF lp] smU have mp: "real m > 0" + from uset_l[OF lp] smU have mp: "real_of_int m > 0" by auto - from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) then show ?thesis using smU by auto @@ -1682,19 +1682,19 @@ assumes lp: "isrlfm p" and nmi: "\ (Ifm (a # bs) (plusinf p))" (is "\ (Ifm (a # bs) (?M p))") and ex: "Ifm (x # bs) p" (is "?I x p") - shows "\(s,m) \ set (uset p). x \ Inum (a#bs) s / real m" - (is "\(s,m) \ ?U p. x \ ?N a s / real m") + shows "\(s,m) \ set (uset p). x \ Inum (a#bs) s / real_of_int m" + (is "\(s,m) \ ?U p. x \ ?N a s / real_of_int m") proof - - have "\(s,m) \ set (uset p). real m * x \ Inum (a#bs) s" - (is "\(s,m) \ ?U p. real m *x \ ?N a s") + have "\(s,m) \ set (uset p). real_of_int m * x \ Inum (a#bs) s" + (is "\(s,m) \ ?U p. real_of_int m *x \ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct) (auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) - then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real m * x \ ?N a s" + then obtain s m where smU: "(s,m) \ set (uset p)" and mx: "real_of_int m * x \ ?N a s" by blast - from uset_l[OF lp] smU have mp: "real m > 0" + from uset_l[OF lp] smU have mp: "real_of_int m > 0" by auto - from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) then show ?thesis using smU by auto @@ -1702,8 +1702,8 @@ lemma lin_dense: assumes lp: "isrlfm p" - and noS: "\t. l < t \ t< u \ t \ (\(t,n). Inum (x#bs) t / real n) ` set (uset p)" - (is "\t. _ \ _ \ t \ (\(t,n). ?N x t / real n ) ` (?U p)") + and noS: "\t. l < t \ t< u \ t \ (\(t,n). Inum (x#bs) t / real_of_int n) ` set (uset p)" + (is "\t. _ \ _ \ t \ (\(t,n). ?N x t / real_of_int n ) ` (?U p)") and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" @@ -1712,163 +1712,163 @@ using lp px noS proof (induct p rule: isrlfm.induct) case (5 c e) - then have cp: "real c > 0" and nb: "numbound0 e" + then have cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all - from 5 have "x * real c + ?N x e < 0" + from 5 have "x * real_of_int c + ?N x e < 0" by (simp add: algebra_simps) - then have pxc: "x < (- ?N x e) / real c" + then have pxc: "x < (- ?N x e) / real_of_int c" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from 5 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real c" + from 5 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto - with ly yu have yne: "y \ - ?N x e / real c" + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto - then consider "y < (-?N x e)/ real c" | "y > (- ?N x e) / real c" + then consider "y < (-?N x e)/ real_of_int c" | "y > (- ?N x e) / real_of_int c" by atomize_elim auto then show ?case proof cases case 1 - then have "y * real c < - ?N x e" + then have "y * real_of_int c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - then have "real c * y + ?N x e < 0" + then have "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) then show ?thesis using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp next case 2 - with yu have eu: "u > (- ?N x e) / real c" + with yu have eu: "u > (- ?N x e) / real_of_int c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" - by (cases "(- ?N x e) / real c > l") auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ l" + by (cases "(- ?N x e) / real_of_int c > l") auto with lx pxc have False by auto then show ?thesis .. qed next case (6 c e) - then have cp: "real c > 0" and nb: "numbound0 e" + then have cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all - from 6 have "x * real c + ?N x e \ 0" + from 6 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) - then have pxc: "x \ (- ?N x e) / real c" + then have pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from 6 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real c" + from 6 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto - with ly yu have yne: "y \ - ?N x e / real c" + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto - then consider "y < (- ?N x e) / real c" | "y > (-?N x e) / real c" + then consider "y < (- ?N x e) / real_of_int c" | "y > (-?N x e) / real_of_int c" by atomize_elim auto then show ?case proof cases case 1 - then have "y * real c < - ?N x e" + then have "y * real_of_int c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - then have "real c * y + ?N x e < 0" + then have "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) then show ?thesis using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp next case 2 - with yu have eu: "u > (- ?N x e) / real c" + with yu have eu: "u > (- ?N x e) / real_of_int c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" - by (cases "(- ?N x e) / real c > l") auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ l" + by (cases "(- ?N x e) / real_of_int c > l") auto with lx pxc have False by auto then show ?thesis .. qed next case (7 c e) - then have cp: "real c > 0" and nb: "numbound0 e" + then have cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all - from 7 have "x * real c + ?N x e > 0" + from 7 have "x * real_of_int c + ?N x e > 0" by (simp add: algebra_simps) - then have pxc: "x > (- ?N x e) / real c" + then have pxc: "x > (- ?N x e) / real_of_int c" by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) - from 7 have noSc: "\t. l < t \ t < u \ t \ (- ?N x e) / real c" + from 7 have noSc: "\t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto - with ly yu have yne: "y \ - ?N x e / real c" + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto - then consider "y > (- ?N x e) / real c" | "y < (-?N x e) / real c" + then consider "y > (- ?N x e) / real_of_int c" | "y < (-?N x e) / real_of_int c" by atomize_elim auto then show ?case proof cases case 1 - then have "y * real c > - ?N x e" + then have "y * real_of_int c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - then have "real c * y + ?N x e > 0" + then have "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) then show ?thesis using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp next case 2 - with ly have eu: "l < (- ?N x e) / real c" + with ly have eu: "l < (- ?N x e) / real_of_int c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" - by (cases "(- ?N x e) / real c > l") auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ u" + by (cases "(- ?N x e) / real_of_int c > l") auto with xu pxc have False by auto then show ?thesis .. qed next case (8 c e) - then have cp: "real c > 0" and nb: "numbound0 e" + then have cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all - from 8 have "x * real c + ?N x e \ 0" + from 8 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) - then have pxc: "x \ (- ?N x e) / real c" + then have pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) - from 8 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real c" + from 8 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto - with ly yu have yne: "y \ - ?N x e / real c" + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto - then consider "y > (- ?N x e) / real c" | "y < (-?N x e) / real c" + then consider "y > (- ?N x e) / real_of_int c" | "y < (-?N x e) / real_of_int c" by atomize_elim auto then show ?case proof cases case 1 - then have "y * real c > - ?N x e" + then have "y * real_of_int c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - then have "real c * y + ?N x e > 0" by (simp add: algebra_simps) + then have "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) then show ?thesis using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp next case 2 - with ly have eu: "l < (- ?N x e) / real c" + with ly have eu: "l < (- ?N x e) / real_of_int c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" - by (cases "(- ?N x e) / real c > l") auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ u" + by (cases "(- ?N x e) / real_of_int c > l") auto with xu pxc have False by auto then show ?thesis .. qed next case (3 c e) - then have cp: "real c > 0" and nb: "numbound0 e" + then have cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all - from cp have cnz: "real c \ 0" + from cp have cnz: "real_of_int c \ 0" by simp - from 3 have "x * real c + ?N x e = 0" + from 3 have "x * real_of_int c + ?N x e = 0" by (simp add: algebra_simps) - then have pxc: "x = (- ?N x e) / real c" + then have pxc: "x = (- ?N x e) / real_of_int c" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) - from 3 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real c" + from 3 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto - with lx xu have yne: "x \ - ?N x e / real c" + with lx xu have yne: "x \ - ?N x e / real_of_int c" by auto with pxc show ?case by simp next case (4 c e) - then have cp: "real c > 0" and nb: "numbound0 e" + then have cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all - from cp have cnz: "real c \ 0" + from cp have cnz: "real_of_int c \ 0" by simp - from 4 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real c" + from 4 have noSc:"\t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto - with ly yu have yne: "y \ - ?N x e / real c" + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto - then have "y* real c \ -?N x e" + then have "y* real_of_int c \ -?N x e" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp - then have "y* real c + ?N x e \ 0" + then have "y* real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) then show ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by (simp add: algebra_simps) @@ -1947,7 +1947,7 @@ and npi: "\ (Ifm (x # bs) (plusinf p))" (is "\ (Ifm (x # bs) (?P p))") and ex: "\x. Ifm (x # bs) p" (is "\x. ?I x p") shows "\(l,n) \ set (uset p). \(s,m) \ set (uset p). - ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" + ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int m) / 2) p" proof - let ?N = "\x t. Inum (x # bs) t" let ?U = "set (uset p)" @@ -1959,22 +1959,22 @@ from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi have npi': "\ (?I a (?P p))" by simp - have "\(l,n) \ set (uset p). \(s,m) \ set (uset p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" + have "\(l,n) \ set (uset p). \(s,m) \ set (uset p). ?I ((?N a l/real_of_int n + ?N a s /real_of_int m) / 2) p" proof - - let ?M = "(\(t,c). ?N a t / real c) ` ?U" + let ?M = "(\(t,c). ?N a t / real_of_int c) ` ?U" have fM: "finite ?M" by auto from rminusinf_uset[OF lp nmi pa] rplusinf_uset[OF lp npi pa] - have "\(l,n) \ set (uset p). \(s,m) \ set (uset p). a \ ?N x l / real n \ a \ ?N x s / real m" + have "\(l,n) \ set (uset p). \(s,m) \ set (uset p). a \ ?N x l / real_of_int n \ a \ ?N x s / real_of_int m" by blast then obtain "t" "n" "s" "m" where tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" - and xs1: "a \ ?N x s / real m" - and tx1: "a \ ?N x t / real n" + and xs1: "a \ ?N x s / real_of_int m" + and tx1: "a \ ?N x t / real_of_int n" by blast from uset_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 - have xs: "a \ ?N a s / real m" and tx: "a \ ?N a t / real n" + have xs: "a \ ?N a s / real_of_int m" and tx: "a \ ?N a t / real_of_int n" by auto from tnU have Mne: "?M \ {}" by auto @@ -1986,19 +1986,19 @@ using fM Mne by simp have uinM: "?u \ ?M" using fM Mne by simp - have tnM: "?N a t / real n \ ?M" + have tnM: "?N a t / real_of_int n \ ?M" using tnU by auto - have smM: "?N a s / real m \ ?M" + have smM: "?N a s / real_of_int m \ ?M" using smU by auto have lM: "\t\ ?M. ?l \ t" using Mne fM by auto have Mu: "\t\ ?M. t \ ?u" using Mne fM by auto - have "?l \ ?N a t / real n" + have "?l \ ?N a t / real_of_int n" using tnM Mne by simp then have lx: "?l \ a" using tx by simp - have "?N a s / real m \ ?u" + have "?N a s / real_of_int m \ ?u" using smM Mne by simp then have xu: "a \ ?u" using xs by simp @@ -2010,13 +2010,13 @@ proof cases case 1 note um = \u \ ?M\ and pu = \?I u p\ - then have "\(tu,nu) \ ?U. u = ?N a tu / real nu" + then have "\(tu,nu) \ ?U. u = ?N a tu / real_of_int nu" by auto - then obtain tu nu where tuU: "(tu, nu) \ ?U" and tuu: "u= ?N a tu / real nu" + then obtain tu nu where tuU: "(tu, nu) \ ?U" and tuu: "u= ?N a tu / real_of_int nu" by blast have "(u + u) / 2 = u" by auto - with pu tuu have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" + with pu tuu have "?I (((?N a tu / real_of_int nu) + (?N a tu / real_of_int nu)) / 2) p" by simp with tuU show ?thesis by blast next @@ -2024,13 +2024,13 @@ note t1M = \t1 \ ?M\ and t2M = \t2\ ?M\ and noM = \\y. t1 < y \ y < t2 \ y \ ?M\ and t1x = \t1 < a\ and xt2 = \a < t2\ and px = \?I a p\ - from t1M have "\(t1u,t1n) \ ?U. t1 = ?N a t1u / real t1n" + from t1M have "\(t1u,t1n) \ ?U. t1 = ?N a t1u / real_of_int t1n" by auto - then obtain t1u t1n where t1uU: "(t1u, t1n) \ ?U" and t1u: "t1 = ?N a t1u / real t1n" + then obtain t1u t1n where t1uU: "(t1u, t1n) \ ?U" and t1u: "t1 = ?N a t1u / real_of_int t1n" by blast - from t2M have "\(t2u,t2n) \ ?U. t2 = ?N a t2u / real t2n" + from t2M have "\(t2u,t2n) \ ?U. t2 = ?N a t2u / real_of_int t2n" by auto - then obtain t2u t2n where t2uU: "(t2u, t2n) \ ?U" and t2u: "t2 = ?N a t2u / real t2n" + then obtain t2u t2n where t2uU: "(t2u, t2n) \ ?U" and t2u: "t2 = ?N a t2u / real_of_int t2n" by blast from t1x xt2 have t1t2: "t1 < t2" by simp @@ -2043,13 +2043,13 @@ qed qed then obtain l n s m where lnU: "(l, n) \ ?U" and smU:"(s, m) \ ?U" - and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" + and pu: "?I ((?N a l / real_of_int n + ?N a s / real_of_int m) / 2) p" by blast from lnU smU uset_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu - have "?I ((?N x l / real n + ?N x s / real m) / 2) p" + have "?I ((?N x l / real_of_int n + ?N x s / real_of_int m) / 2) p" by simp with lnU smU show ?thesis by auto @@ -2063,7 +2063,7 @@ shows "(\x. Ifm (x#bs) p) \ Ifm (x # bs) (minusinf p) \ Ifm (x # bs) (plusinf p) \ (\(t,n) \ set (uset p). \(s,m) \ set (uset p). - Ifm ((((Inum (x # bs) t) / real n + (Inum (x # bs) s) / real m) / 2) # bs) p)" + Ifm ((((Inum (x # bs) t) / real_of_int n + (Inum (x # bs) s) / real_of_int m) / 2) # bs) p)" (is "(\x. ?I x p) \ (?M \ ?P \ ?F)" is "?E = ?D") proof assume px: "\x. ?I x p" @@ -2111,23 +2111,23 @@ then show ?thesis by blast next case 2 - let ?f = "\(t,n). Inum (x # bs) t / real n" + let ?f = "\(t,n). Inum (x # bs) t / real_of_int n" let ?N = "\t. Inum (x # bs) t" { fix t n s m assume "(t, n) \ set (uset p)" and "(s, m) \ set (uset p)" with uset_l[OF lp] have tnb: "numbound0 t" - and np: "real n > 0" and snb: "numbound0 s" and mp: "real m > 0" + and np: "real_of_int n > 0" and snb: "numbound0 s" and mp: "real_of_int m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2 * n * m) > 0" + from np mp have mnp: "real_of_int (2 * n * m) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m) / 2 = ?N ?st / real (2 * n * m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)" using mnp mp np by (simp add: algebra_simps add_divide_distrib) from usubst_I[OF lp mnp st_nb, where x="x" and bs="bs"] - have "?I x (usubst p (?st, 2 * n * m)) = ?I ((?N t / real n + ?N s / real m) / 2) p" + have "?I x (usubst p (?st, 2 * n * m)) = ?I ((?N t / real_of_int n + ?N s / real_of_int m) / 2) p" by (simp only: st[symmetric]) } with rinf_uset[OF lp 2 px] have ?F @@ -2149,11 +2149,11 @@ from rplusinf_ex[OF lp this] show ?thesis . next case 3 - with uset_l[OF lp] have tnb: "numbound0 t" and np: "real k > 0" - and snb: "numbound0 s" and mp: "real l > 0" + with uset_l[OF lp] have tnb: "numbound0 t" and np: "real_of_int k > 0" + and snb: "numbound0 s" and mp: "real_of_int l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" - from np mp have mnp: "real (2 * k * l) > 0" + from np mp have mnp: "real_of_int (2 * k * l) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp @@ -2182,9 +2182,9 @@ lemma uset_cong_aux: assumes Ul: "\(t,n) \ set U. numbound0 t \ n > 0" - shows "((\(t,n). Inum (x # bs) t / real n) ` + shows "((\(t,n). Inum (x # bs) t / real_of_int n) ` (set (map (\((t,n),(s,m)). (Add (Mul m t) (Mul n s), 2 * n * m)) (alluopairs U)))) = - ((\((t,n),(s,m)). (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` (set U \ set U))" + ((\((t,n),(s,m)). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) ` (set U \ set U))" (is "?lhs = ?rhs") proof auto fix t n s m @@ -2197,10 +2197,10 @@ by auto from Ul th have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m) / 2 = ?N ?st / real (2 * n * m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) - then show "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / (2 * real n * real m) - \ (\((t, n), s, m). (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` + then show "(real_of_int m * Inum (x # bs) t + real_of_int n * Inum (x # bs) s) / (2 * real_of_int n * real_of_int m) + \ (\((t, n), s, m). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) ` (set U \ set U)" using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) @@ -2218,10 +2218,10 @@ by auto from Ul tnU have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m) / 2 = ?N ?st / real (2 * n * m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) - let ?P = "\(t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = - (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m') / 2" + let ?P = "\(t',n') (s',m'). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = + (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m') / 2" have Pc:"\a b. ?P a b = ?P b a" by auto from Ul alluopairs_set1 have Up:"\((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" @@ -2235,24 +2235,24 @@ from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto let ?st' = "Add (Mul m' t') (Mul n' s')" - have st': "(?N t' / real n' + ?N s' / real m') / 2 = ?N ?st' / real (2 * n' * m')" + have st': "(?N t' / real_of_int n' + ?N s' / real_of_int m') / 2 = ?N ?st' / real_of_int (2 * n' * m')" using mnz' nnz' by (simp add: algebra_simps add_divide_distrib) - from Pts' have "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 = - (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m') / 2" + from Pts' have "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2 = + (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m') / 2" by simp - also have "\ = (\(t, n). Inum (x # bs) t / real n) + also have "\ = (\(t, n). Inum (x # bs) t / real_of_int n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t', n'), (s', m')))" by (simp add: st') - finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 - \ (\(t, n). Inum (x # bs) t / real n) ` + finally show "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2 + \ (\(t, n). Inum (x # bs) t / real_of_int n) ` (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` set (alluopairs U)" using ts'_U by blast qed lemma uset_cong: assumes lp: "isrlfm p" - and UU': "((\(t,n). Inum (x # bs) t / real n) ` U') = - ((\((t,n),(s,m)). (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` (U \ U))" + and UU': "((\(t,n). Inum (x # bs) t / real_of_int n) ` U') = + ((\((t,n),(s,m)). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) ` (U \ U))" (is "?f ` U' = ?g ` (U \ U)") and U: "\(t,n) \ U. numbound0 t \ n > 0" and U': "\(t,n) \ U'. numbound0 t \ n > 0" @@ -2270,11 +2270,11 @@ and snb: "numbound0 s" and mp: "m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2 * n * m) > 0" - by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) + from np mp have mnp: "real_of_int (2 * n * m) > 0" + by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m) / 2 = ?N ?st / real (2 * n * m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)" using mp np by (simp add: algebra_simps add_divide_distrib) from tnU smU UU' have "?g ((t, n), (s, m)) \ ?f ` U'" by blast @@ -2285,10 +2285,10 @@ done then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t, n), (s, m)) = ?f (t', n')" by blast - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" + from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto from usubst_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp from conjunct1[OF usubst_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] @@ -2316,17 +2316,17 @@ and snb: "numbound0 s" and mp: "m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2 * n * m) > 0" - by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) + from np mp have mnp: "real_of_int (2 * n * m) > 0" + by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m) / 2 = ?N ?st / real (2 * n * m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m) / 2 = ?N ?st / real_of_int (2 * n * m)" using mp np by (simp add: algebra_simps add_divide_distrib) - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" + from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto from usubst_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp with usubst_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?thesis by blast @@ -2348,8 +2348,8 @@ let ?S = "map ?g ?Up" let ?SS = "map simp_num_pair ?S" let ?Y = "remdups ?SS" - let ?f = "\(t,n). ?N t / real n" - let ?h = "\((t,n),(s,m)). (?N t / real n + ?N s / real m) / 2" + let ?f = "\(t,n). ?N t / real_of_int n" + let ?h = "\((t,n),(s,m)). (?N t / real_of_int n + ?N s / real_of_int m) / 2" let ?F = "\p. \a \ set (uset p). \b \ set (uset p). ?I x (usubst p (?g (a, b)))" let ?ep = "evaldjf (simpfm \ (usubst ?q)) ?Y" from rlfm_I[OF simpfm_qf[OF qf]] have lq: "isrlfm ?q" @@ -2403,7 +2403,7 @@ proof - have "bound0 (simpfm (usubst ?q (t, n)))" if tnY: "(t,n) \ set ?Y" for t n proof - - from Y_l that have tnb: "numbound0 t" and np: "real n > 0" + from Y_l that have tnb: "numbound0 t" and np: "real_of_int n > 0" by auto from usubst_I[OF lq np tnb] have "bound0 (usubst ?q (t, n))" by simp @@ -2464,8 +2464,8 @@ val mk_Bound = @{code Bound} o @{code nat_of_integer}; fun num_of_term vs (Free vT) = mk_Bound (find_index (fn vT' => vT = vT') vs) - | num_of_term vs @{term "real (0::int)"} = mk_C 0 - | num_of_term vs @{term "real (1::int)"} = mk_C 1 + | num_of_term vs @{term "real_of_int (0::int)"} = mk_C 0 + | num_of_term vs @{term "real_of_int (1::int)"} = mk_C 1 | num_of_term vs @{term "0::real"} = mk_C 0 | num_of_term vs @{term "1::real"} = mk_C 1 | num_of_term vs (Bound i) = mk_Bound i @@ -2477,7 +2477,7 @@ | num_of_term vs (@{term "op * :: real \ real \ real"} $ t1 $ t2) = (case num_of_term vs t1 of @{code C} i => @{code Mul} (i, num_of_term vs t2) | _ => error "num_of_term: unsupported multiplication") - | num_of_term vs (@{term "real :: int \ real"} $ t') = + | num_of_term vs (@{term "real_of_int :: int \ real"} $ t') = (mk_C (snd (HOLogic.dest_number t')) handle TERM _ => error ("num_of_term: unknown term")) | num_of_term vs t' = @@ -2504,7 +2504,7 @@ @{code A} (fm_of_term (("", dummyT) :: vs) p) | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); -fun term_of_num vs (@{code C} i) = @{term "real :: int \ real"} $ +fun term_of_num vs (@{code C} i) = @{term "real_of_int :: int \ real"} $ HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i) | term_of_num vs (@{code Bound} n) = Free (nth vs (@{code integer_of_nat} n)) | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \ real"} $ term_of_num vs t' diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Tue Nov 10 14:18:41 2015 +0000 @@ -9,7 +9,7 @@ section \Quantifier elimination for @{text "\ (0, 1, +, floor, <)"}\ -declare real_of_int_floor_cancel [simp del] +declare of_int_floor_cancel [simp del] lemma myle: fixes a b :: "'a::{ordered_ab_group_add}" @@ -21,73 +21,51 @@ shows "(a < b) = (0 < b - a)" by (metis le_iff_diff_le_0 less_le_not_le myle) - (* Maybe should be added to the library \ *) -lemma floor_int_eq: "(real n\ x \ x < real (n+1)) = (floor x = n)" -proof( auto) - assume lb: "real n \ x" - and ub: "x < real n + 1" - have "real (floor x) \ x" by simp - hence "real (floor x) < real (n + 1) " using ub by arith - hence "floor x < n+1" by simp - moreover from lb have "n \ floor x" using floor_mono[where x="real n" and y="x"] - by simp ultimately show "floor x = n" by simp -qed - (* Periodicity of dvd *) -lemma dvd_period: - assumes advdd: "(a::int) dvd d" - shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))" - using advdd -proof- - { fix x k - from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"] - have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp} - hence "\x.\k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp - then show ?thesis by simp -qed +lemmas dvd_period = zdvd_period (* The Divisibility relation between reals *) definition rdvd:: "real \ real \ bool" (infixl "rdvd" 50) - where "x rdvd y \ (\k::int. y = x * real k)" + where "x rdvd y \ (\k::int. y = x * real_of_int k)" lemma int_rdvd_real: - "real (i::int) rdvd x = (i dvd (floor x) \ real (floor x) = x)" (is "?l = ?r") + "real_of_int (i::int) rdvd x = (i dvd (floor x) \ real_of_int (floor x) = x)" (is "?l = ?r") proof assume "?l" - hence th: "\ k. x=real (i*k)" by (simp add: rdvd_def) - hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult) - with th have "\ k. real (floor x) = real (i*k)" by simp - hence "\ k. floor x = i*k" by (simp only: real_of_int_inject) + hence th: "\ k. x=real_of_int (i*k)" by (simp add: rdvd_def) + hence th': "real_of_int (floor x) = x" by (auto simp del: of_int_mult) + with th have "\ k. real_of_int (floor x) = real_of_int (i*k)" by simp + hence "\ k. floor x = i*k" by blast thus ?r using th' by (simp add: dvd_def) next assume "?r" hence "(i::int) dvd \x::real\" .. - hence "\ k. real (floor x) = real (i*k)" - by (simp only: real_of_int_inject) (simp add: dvd_def) + hence "\ k. real_of_int (floor x) = real_of_int (i*k)" + using dvd_def by blast thus ?l using \?r\ by (simp add: rdvd_def) qed -lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)" - by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only: real_of_int_mult[symmetric]) - - -lemma rdvd_abs1: "(abs (real d) rdvd t) = (real (d ::int) rdvd t)" +lemma int_rdvd_iff: "(real_of_int (i::int) rdvd real_of_int t) = (i dvd t)" + by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only: of_int_mult[symmetric]) + + +lemma rdvd_abs1: "(abs (real_of_int d) rdvd t) = (real_of_int (d ::int) rdvd t)" proof - assume d: "real d rdvd t" - from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" + assume d: "real_of_int d rdvd t" + from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real_of_int (floor t) = t" by auto from iffD2[OF abs_dvd_iff] d2 have "(abs d) dvd (floor t)" by blast - with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast - thus "abs (real d) rdvd t" by simp + with ti int_rdvd_real[symmetric] have "real_of_int (abs d) rdvd t" by blast + thus "abs (real_of_int d) rdvd t" by simp next - assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp - with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real (floor t) =t" + assume "abs (real_of_int d) rdvd t" hence "real_of_int (abs d) rdvd t" by simp + with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real_of_int (floor t) =t" by auto from iffD1[OF abs_dvd_iff] d2 have "d dvd floor t" by blast - with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast + with ti int_rdvd_real[symmetric] show "real_of_int d rdvd t" by blast qed -lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)" +lemma rdvd_minus: "(real_of_int (d::int) rdvd t) = (real_of_int d rdvd -t)" apply (auto simp add: rdvd_def) apply (rule_tac x="-k" in exI, simp) apply (rule_tac x="-k" in exI, simp) @@ -98,7 +76,7 @@ lemma rdvd_mult: assumes knz: "k\0" - shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)" + shows "(real_of_int (n::int) * real_of_int (k::int) rdvd x * real_of_int k) = (real_of_int n rdvd x)" using knz by (simp add: rdvd_def) (*********************************************************************************) @@ -122,18 +100,18 @@ (* Semantics of numeral terms (num) *) primrec Inum :: "real list \ num \ real" where - "Inum bs (C c) = (real c)" + "Inum bs (C c) = (real_of_int c)" | "Inum bs (Bound n) = bs!n" -| "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" +| "Inum bs (CN n c a) = (real_of_int c) * (bs!n) + (Inum bs a)" | "Inum bs (Neg a) = -(Inum bs a)" | "Inum bs (Add a b) = Inum bs a + Inum bs b" | "Inum bs (Sub a b) = Inum bs a - Inum bs b" -| "Inum bs (Mul c a) = (real c) * Inum bs a" -| "Inum bs (Floor a) = real (floor (Inum bs a))" -| "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b" -definition "isint t bs \ real (floor (Inum bs t)) = Inum bs t" - -lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)" +| "Inum bs (Mul c a) = (real_of_int c) * Inum bs a" +| "Inum bs (Floor a) = real_of_int (floor (Inum bs a))" +| "Inum bs (CF c a b) = real_of_int c * real_of_int (floor (Inum bs a)) + Inum bs b" +definition "isint t bs \ real_of_int (floor (Inum bs t)) = Inum bs t" + +lemma isint_iff: "isint n bs = (real_of_int (floor (Inum bs n)) = Inum bs n)" by (simp add: isint_def) lemma isint_Floor: "isint (Floor n) bs" @@ -143,10 +121,10 @@ proof- let ?e = "Inum bs e" let ?fe = "floor ?e" - assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff) - have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp - also have "\ = real (c* ?fe)" by (simp only: floor_real_of_int) - also have "\ = real c * ?e" using efe by simp + assume be: "isint e bs" hence efe:"real_of_int ?fe = ?e" by (simp add: isint_iff) + have "real_of_int ((floor (Inum bs (Mul c e)))) = real_of_int (floor (real_of_int (c * ?fe)))" using efe by simp + also have "\ = real_of_int (c* ?fe)" using floor_of_int by blast + also have "\ = real_of_int c * ?e" using efe by simp finally show ?thesis using isint_iff by simp qed @@ -154,9 +132,9 @@ proof- let ?I = "\ t. Inum bs t" assume ie: "isint e bs" - hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) - have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th) - also have "\ = - real (floor (?I e))" by simp + hence th: "real_of_int (floor (?I e)) = ?I e" by (simp add: isint_def) + have "real_of_int (floor (?I (Neg e))) = real_of_int (floor (- (real_of_int (floor (?I e)))))" by (simp add: th) + also have "\ = - real_of_int (floor (?I e))" by simp finally show "isint (Neg e) bs" by (simp add: isint_def th) qed @@ -164,9 +142,9 @@ assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs" proof- let ?I = "\ t. Inum bs t" - from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) - have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th) - also have "\ = real (c- floor (?I e))" by simp + from ie have th: "real_of_int (floor (?I e)) = ?I e" by (simp add: isint_def) + have "real_of_int (floor (?I (Sub (C c) e))) = real_of_int (floor ((real_of_int (c -floor (?I e)))))" by (simp add: th) + also have "\ = real_of_int (c- floor (?I e))" by simp finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th) qed @@ -176,9 +154,9 @@ proof- let ?a = "Inum bs a" let ?b = "Inum bs b" - from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" + from ai bi isint_iff have "real_of_int (floor (?a + ?b)) = real_of_int (floor (real_of_int (floor ?a) + real_of_int (floor ?b)))" by simp - also have "\ = real (floor ?a) + real (floor ?b)" by simp + also have "\ = real_of_int (floor ?a) + real_of_int (floor ?b)" by simp also have "\ = ?a + ?b" using ai bi isint_iff by simp finally show "isint (Add a b) bs" by (simp add: isint_iff) qed @@ -219,8 +197,8 @@ | "Ifm bs (Ge a) = (Inum bs a \ 0)" | "Ifm bs (Eq a) = (Inum bs a = 0)" | "Ifm bs (NEq a) = (Inum bs a \ 0)" -| "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)" -| "Ifm bs (NDvd i b) = (\(real i rdvd Inum bs b))" +| "Ifm bs (Dvd i b) = (real_of_int i rdvd Inum bs b)" +| "Ifm bs (NDvd i b) = (\(real_of_int i rdvd Inum bs b))" | "Ifm bs (NOT p) = (\ (Ifm bs p))" | "Ifm bs (And p q) = (Ifm bs p \ Ifm bs q)" | "Ifm bs (Or p q) = (Ifm bs p \ Ifm bs q)" @@ -607,7 +585,7 @@ lemma reducecoeffh: assumes gt: "dvdnumcoeff t g" and gp: "g > 0" - shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" + shows "real_of_int g *(Inum bs (reducecoeffh t g)) = Inum bs t" using gt proof(induct t rule: reducecoeffh.induct) case (1 i) hence gd: "g dvd i" by simp @@ -708,7 +686,7 @@ from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . qed -lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" +lemma reducecoeff: "real_of_int (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" proof- let ?g = "numgcd t" have "?g \ 0" by (simp add: numgcd_pos) @@ -757,7 +735,7 @@ apply (case_tac "lex_bnd t1 t2", simp_all) apply (case_tac "c1+c2 = 0") apply (case_tac "t1 = t2") - apply (simp_all add: algebra_simps distrib_right[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add distrib_right) + apply (simp_all add: algebra_simps distrib_right[symmetric] of_int_mult[symmetric] of_int_add[symmetric]del: of_int_mult of_int_add distrib_right) done lemma numadd_nb[simp]: "\ numbound0 t ; numbound0 s\ \ numbound0 (numadd (t,s))" @@ -852,15 +830,15 @@ hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" by (cases ?tv) (auto simp add: numfloor_def Let_def split_def) from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ - hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp - also have "\ = real (floor (?N ?tv) + (floor (?N ?ti)))" + hence "?N (Floor t) = real_of_int (floor (?N (Add ?tv ?ti)))" by simp + also have "\ = real_of_int (floor (?N ?tv) + (floor (?N ?ti)))" by (simp,subst tii[simplified isint_iff, symmetric]) simp also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) finally have ?thesis using th1 by simp} moreover {fix v assume H:"?tv = C v" from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ - hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp - also have "\ = real (floor (?N ?tv) + (floor (?N ?ti)))" + hence "?N (Floor t) = real_of_int (floor (?N (Add ?tv ?ti)))" by simp + also have "\ = real_of_int (floor (?N ?tv) + (floor (?N ?ti)))" by (simp,subst tii[simplified isint_iff, symmetric]) simp also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) finally have ?thesis by (simp add: H numfloor_def Let_def split_def) } @@ -951,7 +929,7 @@ else (t',n))))" lemma simp_num_pair_ci: - shows "((\ (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real n) (t,n))" + shows "((\ (t,n). Inum bs t / real_of_int n) (simp_num_pair (t,n))) = ((\ (t,n). Inum bs t / real_of_int n) (t,n))" (is "?lhs = ?rhs") proof- let ?t' = "simpnum t" @@ -975,12 +953,12 @@ have gpdd: "?g' dvd n" by simp have gpdgp: "?g' dvd ?g'" by simp from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] - have th2:"real ?g' * ?t = Inum bs ?t'" by simp - from nnz g1 g'1 have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) - also have "\ = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp - also have "\ = (Inum bs ?t' / real n)" + have th2:"real_of_int ?g' * ?t = Inum bs ?t'" by simp + from nnz g1 g'1 have "?lhs = ?t / real_of_int (n div ?g')" by (simp add: simp_num_pair_def Let_def) + also have "\ = (real_of_int ?g' * ?t) / (real_of_int ?g' * (real_of_int (n div ?g')))" by simp + also have "\ = (Inum bs ?t' / real_of_int n)" using real_of_int_div[OF gpdd] th2 gp0 by simp - finally have "?lhs = Inum bs t / real n" by simp + finally have "?lhs = Inum bs t / real_of_int n" by simp then have ?thesis using nnz g1 g'1 by (simp add: simp_num_pair_def) } ultimately have ?thesis by blast } ultimately have ?thesis by blast } @@ -1092,27 +1070,27 @@ lemma check_int: "check_int t \ isint t bs" by (induct t) (auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF) -lemma rdvd_left1_int: "real \t\ = t \ 1 rdvd t" +lemma rdvd_left1_int: "real_of_int \t\ = t \ 1 rdvd t" by (simp add: rdvd_def,rule_tac x="\t\" in exI) simp lemma rdvd_reduce: assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0" - shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)" + shows "real_of_int (d::int) rdvd real_of_int (c::int)*t = (real_of_int (d div g) rdvd real_of_int (c div g)*t)" proof - assume d: "real d rdvd real c * t" - from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto + assume d: "real_of_int d rdvd real_of_int c * t" + from d rdvd_def obtain k where k_def: "real_of_int c * t = real_of_int d* real_of_int (k::int)" by auto from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto - from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp - hence "real kc * t = real kd * real k" using gp by simp - hence th:"real kd rdvd real kc * t" using rdvd_def by blast + from k_def kd_def kc_def have "real_of_int g * real_of_int kc * t = real_of_int g * real_of_int kd * real_of_int k" by simp + hence "real_of_int kc * t = real_of_int kd * real_of_int k" using gp by simp + hence th:"real_of_int kd rdvd real_of_int kc * t" using rdvd_def by blast from kd_def gp have th':"kd = d div g" by simp from kc_def gp have "kc = c div g" by simp - with th th' show "real (d div g) rdvd real (c div g) * t" by simp + with th th' show "real_of_int (d div g) rdvd real_of_int (c div g) * t" by simp next - assume d: "real (d div g) rdvd real (c div g) * t" + assume d: "real_of_int (d div g) rdvd real_of_int (c div g) * t" from gp have gnz: "g \ 0" by simp - thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gd] real_of_int_div[OF gc] by simp + thus "real_of_int d rdvd real_of_int c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real_of_int (c div g) * t"] real_of_int_div[OF gd] real_of_int_div[OF gc] by simp qed definition simpdvd :: "int \ num \ (int \ num)" where @@ -1143,11 +1121,11 @@ have gpdd: "?g' dvd d" by simp have gpdgp: "?g' dvd ?g'" by simp from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] - have th2:"real ?g' * ?t = Inum bs t" by simp + have th2:"real_of_int ?g' * ?t = Inum bs t" by simp from assms g1 g0 g'1 have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)" by (simp add: simpdvd_def Let_def) - also have "\ = (real d rdvd (Inum bs t))" + also have "\ = (real_of_int d rdvd (Inum bs t))" using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified div_self[OF gp0]] th2[symmetric] by simp finally have ?thesis by simp } @@ -1190,9 +1168,9 @@ have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp + hence gp: "real_of_int ?g > 0" by simp + have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a < 0 = (real_of_int ?g * ?r < real_of_int ?g * 0)" by simp also have "\ = (?r < 0)" using gp by (simp only: mult_less_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} @@ -1207,9 +1185,9 @@ have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a \ 0 = (real ?g * ?r \ real ?g * 0)" by simp + hence gp: "real_of_int ?g > 0" by simp + have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a \ 0 = (real_of_int ?g * ?r \ real_of_int ?g * 0)" by simp also have "\ = (?r \ 0)" using gp by (simp only: mult_le_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} @@ -1224,9 +1202,9 @@ have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp + hence gp: "real_of_int ?g > 0" by simp + have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a > 0 = (real_of_int ?g * ?r > real_of_int ?g * 0)" by simp also have "\ = (?r > 0)" using gp by (simp only: mult_less_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} @@ -1241,9 +1219,9 @@ have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a \ 0 = (real ?g * ?r \ real ?g * 0)" by simp + hence gp: "real_of_int ?g > 0" by simp + have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a \ 0 = (real_of_int ?g * ?r \ real_of_int ?g * 0)" by simp also have "\ = (?r \ 0)" using gp by (simp only: mult_le_cancel_left) simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} @@ -1258,9 +1236,9 @@ have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp + hence gp: "real_of_int ?g > 0" by simp + have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a = 0 = (real_of_int ?g * ?r = 0)" by simp also have "\ = (?r = 0)" using gp by simp finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} @@ -1275,9 +1253,9 @@ have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) - hence gp: "real ?g > 0" by simp - have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) - with sa have "Inum bs a \ 0 = (real ?g * ?r \ 0)" by simp + hence gp: "real_of_int ?g > 0" by simp + have "Inum bs ?sa = real_of_int ?g* ?r" by (simp add: reducecoeff) + with sa have "Inum bs a \ 0 = (real_of_int ?g * ?r \ 0)" by simp also have "\ = (?r \ 0)" using gp by simp finally have ?case using H by (cases "?sa") (simp_all add: Let_def) } @@ -1471,7 +1449,7 @@ termination by (relation "measure num_size") auto lemma zsplit0_I: - shows "\ n a. zsplit0 t = (n,a) \ (Inum ((real (x::int)) #bs) (CN 0 n a) = Inum (real x #bs) t) \ numbound0 a" + shows "\ n a. zsplit0 t = (n,a) \ (Inum ((real_of_int (x::int)) #bs) (CN 0 n a) = Inum (real_of_int x #bs) t) \ numbound0 a" (is "\ n a. ?S t = (n,a) \ (?I x (CN 0 n a) = ?I x t) \ ?N a") proof(induct t rule: zsplit0.induct) case (1 c n a) thus ?case by auto @@ -1500,7 +1478,7 @@ ultimately have th: "a=Add ?as ?at \ n=?ns + ?nt" using 6 by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast - from 6 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \ numbound0 xb)" by blast (*FIXME*) + from 6 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int x # bs) t \ numbound0 xb)" by blast (*FIXME*) with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from abjs 6 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast from th3[simplified] th2[simplified] th[simplified] show ?case @@ -1516,7 +1494,7 @@ ultimately have th: "a=Sub ?as ?at \ n=?ns - ?nt" using 7 by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast - from 7 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \ numbound0 xb)" by blast (*FIXME*) + from 7 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int x # bs) t \ numbound0 xb)" by blast (*FIXME*) with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from abjs 7 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast from th3[simplified] th2[simplified] th[simplified] show ?case @@ -1528,7 +1506,7 @@ have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \ n=i*?nt" using 8 by (simp add: Let_def split_def) from abj 8 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast - hence "?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp + hence "?I x (Mul i t) = (real_of_int i) * ?I x (CN 0 ?nt ?at)" by simp also have "\ = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: distrib_left) finally show ?case using th th2 by simp next @@ -1539,13 +1517,13 @@ by (simp add: Let_def split_def) from abj 9 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast hence na: "?N a" using th by simp - have th': "(real ?nt)*(real x) = real (?nt * x)" by simp + have th': "(real_of_int ?nt)*(real_of_int x) = real_of_int (?nt * x)" by simp have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp - also have "\ = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp - also have "\ = real (floor (?I x ?at + real (?nt* x)))" by (simp add: ac_simps) - also have "\ = real (floor (?I x ?at) + (?nt* x))" - using floor_add[where x="?I x ?at" and a="?nt* x"] by simp - also have "\ = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: ac_simps) + also have "\ = real_of_int (floor ((real_of_int ?nt)* real_of_int(x) + ?I x ?at))" by simp + also have "\ = real_of_int (floor (?I x ?at + real_of_int (?nt* x)))" by (simp add: ac_simps) + also have "\ = real_of_int (floor (?I x ?at) + (?nt* x))" + using floor_add_of_int[of "?I x ?at" "?nt* x"] by simp + also have "\ = real_of_int (?nt)*(real_of_int x) + real_of_int (floor (?I x ?at))" by (simp add: ac_simps) finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp with na show ?case by simp qed @@ -1643,72 +1621,72 @@ "zlfm p = p" (hints simp add: fmsize_pos) lemma split_int_less_real: - "(real (a::int) < b) = (a < floor b \ (a = floor b \ real (floor b) < b))" + "(real_of_int (a::int) < b) = (a < floor b \ (a = floor b \ real_of_int (floor b) < b))" proof( auto) - assume alb: "real a < b" and agb: "\ a < floor b" - from agb have "floor b \ a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq) + assume alb: "real_of_int a < b" and agb: "\ a < floor b" + from agb have "floor b \ a" by simp hence th: "b < real_of_int a + 1" by (simp only: floor_le_iff) from floor_eq[OF alb th] show "a= floor b" by simp next assume alb: "a < floor b" - hence "real a < real (floor b)" by simp - moreover have "real (floor b) \ b" by simp ultimately show "real a < b" by arith + hence "real_of_int a < real_of_int (floor b)" by simp + moreover have "real_of_int (floor b) \ b" by simp ultimately show "real_of_int a < b" by arith qed lemma split_int_less_real': - "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \ (real a - real (floor (-b)) = 0 \ real (floor (-b)) + b < 0))" + "(real_of_int (a::int) + b < 0) = (real_of_int a - real_of_int (floor(-b)) < 0 \ (real_of_int a - real_of_int (floor (-b)) = 0 \ real_of_int (floor (-b)) + b < 0))" proof- - have "(real a + b <0) = (real a < -b)" by arith + have "(real_of_int a + b <0) = (real_of_int a < -b)" by arith with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith qed lemma split_int_gt_real': - "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \ (real a + real (floor b) = 0 \ real (floor b) - b < 0))" + "(real_of_int (a::int) + b > 0) = (real_of_int a + real_of_int (floor b) > 0 \ (real_of_int a + real_of_int (floor b) = 0 \ real_of_int (floor b) - b < 0))" proof- - have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith - show ?thesis using myless[of _ "real (floor b)"] + have th: "(real_of_int a + b >0) = (real_of_int (-a) + (-b)< 0)" by arith + show ?thesis using myless[of _ "real_of_int (floor b)"] by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) (simp add: algebra_simps,arith) qed lemma split_int_le_real: - "(real (a::int) \ b) = (a \ floor b \ (a = floor b \ real (floor b) < b))" + "(real_of_int (a::int) \ b) = (a \ floor b \ (a = floor b \ real_of_int (floor b) < b))" proof( auto) - assume alb: "real a \ b" and agb: "\ a \ floor b" - from alb have "floor (real a) \ floor b " by (simp only: floor_mono) + assume alb: "real_of_int a \ b" and agb: "\ a \ floor b" + from alb have "floor (real_of_int a) \ floor b " by (simp only: floor_mono) hence "a \ floor b" by simp with agb show "False" by simp next assume alb: "a \ floor b" - hence "real a \ real (floor b)" by (simp only: floor_mono) - also have "\\ b" by simp finally show "real a \ b" . + hence "real_of_int a \ real_of_int (floor b)" by (simp only: floor_mono) + also have "\\ b" by simp finally show "real_of_int a \ b" . qed lemma split_int_le_real': - "(real (a::int) + b \ 0) = (real a - real (floor(-b)) \ 0 \ (real a - real (floor (-b)) = 0 \ real (floor (-b)) + b < 0))" + "(real_of_int (a::int) + b \ 0) = (real_of_int a - real_of_int (floor(-b)) \ 0 \ (real_of_int a - real_of_int (floor (-b)) = 0 \ real_of_int (floor (-b)) + b < 0))" proof- - have "(real a + b \0) = (real a \ -b)" by arith + have "(real_of_int a + b \0) = (real_of_int a \ -b)" by arith with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith qed lemma split_int_ge_real': - "(real (a::int) + b \ 0) = (real a + real (floor b) \ 0 \ (real a + real (floor b) = 0 \ real (floor b) - b < 0))" + "(real_of_int (a::int) + b \ 0) = (real_of_int a + real_of_int (floor b) \ 0 \ (real_of_int a + real_of_int (floor b) = 0 \ real_of_int (floor b) - b < 0))" proof- - have th: "(real a + b \0) = (real (-a) + (-b) \ 0)" by arith + have th: "(real_of_int a + b \0) = (real_of_int (-a) + (-b) \ 0)" by arith show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"]) (simp add: algebra_simps ,arith) qed -lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \ b = real (floor b))" (is "?l = ?r") +lemma split_int_eq_real: "(real_of_int (a::int) = b) = ( a = floor b \ b = real_of_int (floor b))" (is "?l = ?r") by auto -lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \ real (floor (-b)) + b = 0)" (is "?l = ?r") +lemma split_int_eq_real': "(real_of_int (a::int) + b = 0) = ( a - floor (-b) = 0 \ real_of_int (floor (-b)) + b = 0)" (is "?l = ?r") proof- - have "?l = (real a = -b)" by arith + have "?l = (real_of_int a = -b)" by arith with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith qed lemma zlfm_I: assumes qfp: "qfree p" - shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \ iszlfm (zlfm p) (real (i::int) #bs)" + shows "(Ifm (real_of_int i #bs) (zlfm p) = Ifm (real_of_int i# bs) p) \ iszlfm (zlfm p) (real_of_int (i::int) #bs)" (is "(?I (?l p) = ?I p) \ ?L (?l p)") using qfp proof(induct p rule: zlfm.induct) @@ -1717,8 +1695,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] @@ -1726,13 +1704,13 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) + have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) + have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Lt a)))" by (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast @@ -1742,8 +1720,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] @@ -1751,13 +1729,13 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Le a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + have "?I (Le a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Le a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + have "?I (Le a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast @@ -1767,8 +1745,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] @@ -1776,13 +1754,13 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) + have "?I (Gt a) = (real_of_int (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) + have "?I (Gt a) = (real_of_int (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast @@ -1792,8 +1770,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] @@ -1801,13 +1779,13 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Ge a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + have "?I (Ge a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Ge a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + have "?I (Ge a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast @@ -1817,8 +1795,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] @@ -1826,14 +1804,14 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (Eq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult) + have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (Eq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith) + have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult,arith) finally have ?case using l by simp} ultimately show ?case by blast next @@ -1842,8 +1820,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] @@ -1851,14 +1829,14 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NEq a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also have "\ = (?I (?l (NEq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult) + have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also have "\ = (?I (?l (NEq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult) finally have ?case using l by simp} moreover {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NEq a) = (real (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) - also from cn cnz have "\ = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith) + have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) + also from cn cnz have "\ = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia of_int_mult[symmetric] del: of_int_mult,arith) finally have ?case using l by simp} ultimately show ?case by blast next @@ -1867,8 +1845,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith moreover { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) @@ -1880,31 +1858,31 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" + have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" using Ia by (simp add: Let_def split_def) - also have "\ = (real (abs j) rdvd real (?c*i) + (?N ?r))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps) + also have "\ = (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r))" + by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp + also have "\ = ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (Dvd j a)))" using cp cnz jnz by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: ac_simps) + del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by simp } moreover {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" + have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" using Ia by (simp add: Let_def split_def) - also have "\ = (real (abs j) rdvd real (?c*i) + (?N ?r))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps) + also have "\ = (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r))" + by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp + also have "\ = ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (Dvd j a)))" using cn cnz jnz - using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] + using rdvd_minus [where d="abs j" and t="real_of_int (?c*i + floor (?N ?r))", simplified, symmetric] by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: ac_simps) + del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by blast } ultimately show ?case by blast next @@ -1913,8 +1891,8 @@ let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto - let ?N = "\ t. Inum (real i#bs) t" + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + let ?N = "\ t. Inum (real_of_int i#bs) t" have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith moreover {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) @@ -1926,31 +1904,31 @@ moreover {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NDvd j a) = (\ (real j rdvd (real (?c * i) + (?N ?r))))" + have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" using Ia by (simp add: Let_def split_def) - also have "\ = (\ (real (abs j) rdvd real (?c*i) + (?N ?r)))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps) + also have "\ = (\ (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r)))" + by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp + also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r)))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (NDvd j a)))" using cp cnz jnz by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: ac_simps) + del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by simp } moreover {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NDvd j a) = (\ (real j rdvd (real (?c * i) + (?N ?r))))" + have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" using Ia by (simp add: Let_def split_def) - also have "\ = (\ (real (abs j) rdvd real (?c*i) + (?N ?r)))" - by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp - also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \ - (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" - by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: ac_simps) + also have "\ = (\ (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r)))" + by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp + also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r)))))" + by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (NDvd j a)))" using cn cnz jnz - using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] + using rdvd_minus [where d="abs j" and t="real_of_int (?c*i + floor (?N ?r))", simplified, symmetric] by (simp add: Let_def split_def int_rdvd_iff[symmetric] - del: real_of_int_mult) (auto simp add: ac_simps) + del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by blast } ultimately show ?case by blast qed auto @@ -2040,7 +2018,7 @@ lemma minusinf_inf: assumes linp: "iszlfm p (a # bs)" - shows "\ (z::int). \ x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p" + shows "\ (z::int). \ x < z. Ifm ((real_of_int x)#bs) (minusinf p) = Ifm ((real_of_int x)#bs) p" (is "?P p" is "\ (z::int). \ x < z. ?I x (?M p) = ?I x p") using linp proof (induct p rule: minusinf.induct) @@ -2064,173 +2042,173 @@ next case (3 c e) then have "c > 0" by simp - hence rcpos: "real c > 0" by simp + hence rcpos: "real_of_int c > 0" by simp from 3 have nbe: "numbound0 e" by simp fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) + have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))" + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int - assume A: "real x + 1 \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" + hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp + with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - hence "real c * real x + Inum (y # bs) e \ 0"using rcpos by simp - thus "real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp + hence "real_of_int c * real_of_int x + Inum (y # bs) e \ 0"using rcpos by simp + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] by simp qed thus ?case by blast next case (4 c e) - then have "c > 0" by simp hence rcpos: "real c > 0" by simp + then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 4 have nbe: "numbound0 e" by simp fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) + have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))" + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int - assume A: "real x + 1 \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" + hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp + with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - hence "real c * real x + Inum (y # bs) e \ 0"using rcpos by simp - thus "real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp + hence "real_of_int c * real_of_int x + Inum (y # bs) e \ 0"using rcpos by simp + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] by simp qed thus ?case by blast next case (5 c e) - then have "c > 0" by simp hence rcpos: "real c > 0" by simp + then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 5 have nbe: "numbound0 e" by simp fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) + have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))" + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int - assume A: "real x + 1 \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" + hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp + with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "real c * real x + Inum (real x # bs) e < 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next case (6 c e) - then have "c > 0" by simp hence rcpos: "real c > 0" by simp + then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 6 have nbe: "numbound0 e" by simp fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) + have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))" + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int - assume A: "real x + 1 \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" + hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp + with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next case (7 c e) - then have "c > 0" by simp hence rcpos: "real c > 0" by simp + then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 7 have nbe: "numbound0 e" by simp fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) + have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))" + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int - assume A: "real x + 1 \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" + hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp + with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "\ (real c * real x + Inum (real x # bs) e>0)" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + thus "\ (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e>0)" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next case (8 c e) - then have "c > 0" by simp hence rcpos: "real c > 0" by simp + then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 8 have nbe: "numbound0 e" by simp fix y - have "\ x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))" - proof (simp add: less_floor_eq , rule allI, rule impI) + have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))" + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int - assume A: "real x + 1 \ - (Inum (y # bs) e / real c)" - hence th1:"real x < - (Inum (y # bs) e / real c)" by simp - with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" + assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" + hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp + with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "\ real c * real x + Inum (real x # bs) e \ 0" - using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp + thus "\ real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast qed simp_all lemma minusinf_repeats: assumes d: "d_\ p d" and linp: "iszlfm p (a # bs)" - shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)" + shows "Ifm ((real_of_int(x - k*d))#bs) (minusinf p) = Ifm (real_of_int x #bs) (minusinf p)" using linp d proof(induct p rule: iszlfm.induct) case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ hence "\ k. d=i*k" by (simp add: dvd_def) then obtain "di" where di_def: "d=i*di" by blast show ?case - proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI) + proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI) assume - "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" + "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") - hence "\ (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" + hence "\ (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" by (simp add: algebra_simps di_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))" by (simp add: algebra_simps) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast - thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast + thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" using rdvd_def by simp next assume - "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") - hence "\ (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" + "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e") + hence "\ (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)" by simp + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int i * real_of_int di)" by (simp add: di_def) + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int (l - c*k*di))" by (simp add: algebra_simps) + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l)" by blast - thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp + thus "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" using rdvd_def by simp qed next case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ hence "\ k. d=i*k" by (simp add: dvd_def) then obtain "di" where di_def: "d=i*di" by blast show ?case - proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI) + proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI) assume - "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" + "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") - hence "\ (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" + hence "\ (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def) + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" by (simp add: algebra_simps di_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))" by (simp add: algebra_simps) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast - thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast + thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" using rdvd_def by simp next assume - "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") - hence "\ (l::int). ?rc*?rx+?e = ?ri * (real l)" + "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e") + hence "\ (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)" by simp - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int i * real_of_int di)" by (simp add: di_def) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int (l - c*k*di))" by (simp add: algebra_simps) - hence "\ (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" + hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l)" by blast - thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" + thus "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" using rdvd_def by simp qed -qed (auto simp add: numbound0_I[where bs="bs" and b="real(x - k*d)" and b'="real x"] simp del: real_of_int_mult real_of_int_diff) +qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int(x - k*d)" and b'="real_of_int x"] simp del: of_int_mult of_int_diff) lemma minusinf_ex: - assumes lin: "iszlfm p (real (a::int) #bs)" - and exmi: "\ (x::int). Ifm (real x#bs) (minusinf p)" (is "\ x. ?P1 x") - shows "\ (x::int). Ifm (real x#bs) p" (is "\ x. ?P x") + assumes lin: "iszlfm p (real_of_int (a::int) #bs)" + and exmi: "\ (x::int). Ifm (real_of_int x#bs) (minusinf p)" (is "\ x. ?P1 x") + shows "\ (x::int). Ifm (real_of_int x#bs) p" (is "\ x. ?P x") proof- let ?d = "\ p" from \ [OF lin] have dpos: "?d >0" by simp @@ -2241,9 +2219,9 @@ qed lemma minusinf_bex: - assumes lin: "iszlfm p (real (a::int) #bs)" - shows "(\ (x::int). Ifm (real x#bs) (minusinf p)) = - (\ (x::int)\ {1..\ p}. Ifm (real x#bs) (minusinf p))" + assumes lin: "iszlfm p (real_of_int (a::int) #bs)" + shows "(\ (x::int). Ifm (real_of_int x#bs) (minusinf p)) = + (\ (x::int)\ {1..\ p}. Ifm (real_of_int x#bs) (minusinf p))" (is "(\ x. ?P x) = _") proof- let ?d = "\ p" @@ -2338,30 +2316,30 @@ lemma mirror_\_\: assumes lp: "iszlfm p (a#bs)" - shows "(Inum (real (i::int)#bs)) ` set (\ p) = (Inum (real i#bs)) ` set (\ (mirror p))" + shows "(Inum (real_of_int (i::int)#bs)) ` set (\ p) = (Inum (real_of_int i#bs)) ` set (\ (mirror p))" using lp by (induct p rule: mirror.induct) auto lemma mirror: assumes lp: "iszlfm p (a#bs)" - shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" + shows "Ifm (real_of_int (x::int)#bs) (mirror p) = Ifm (real_of_int (- x)#bs) p" using lp proof(induct p rule: iszlfm.induct) case (9 j c e) - have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = - (real j rdvd - (real c * real x - Inum (real x # bs) e))" + have th: "(real_of_int j rdvd real_of_int c * real_of_int x - Inum (real_of_int x # bs) e) = + (real_of_int j rdvd - (real_of_int c * real_of_int x - Inum (real_of_int x # bs) e))" by (simp only: rdvd_minus[symmetric]) from 9 th show ?case by (simp add: algebra_simps - numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) + numbound0_I[where bs="bs" and b'="real_of_int x" and b="- real_of_int x"]) next case (10 j c e) - have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = - (real j rdvd - (real c * real x - Inum (real x # bs) e))" + have th: "(real_of_int j rdvd real_of_int c * real_of_int x - Inum (real_of_int x # bs) e) = + (real_of_int j rdvd - (real_of_int c * real_of_int x - Inum (real_of_int x # bs) e))" by (simp only: rdvd_minus[symmetric]) from 10 th show ?case by (simp add: algebra_simps - numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) -qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"]) + numbound0_I[where bs="bs" and b'="real_of_int x" and b="- real_of_int x"]) +qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int x" and b'="- real_of_int x"]) lemma mirror_l: "iszlfm p (a#bs) \ iszlfm (mirror p) (a#bs)" by (induct p rule: mirror.induct) (auto simp add: isint_neg) @@ -2375,8 +2353,8 @@ lemma mirror_ex: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "(\ (x::int). Ifm (real x#bs) (mirror p)) = (\ (x::int). Ifm (real x#bs) p)" + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" + shows "(\ (x::int). Ifm (real_of_int x#bs) (mirror p)) = (\ (x::int). Ifm (real_of_int x#bs) p)" (is "(\ x. ?I x ?mp) = (\ x. ?I x p)") proof(auto) fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast @@ -2425,7 +2403,7 @@ qed (auto simp add: lcm_pos_int) lemma a_\: assumes linp: "iszlfm p (a #bs)" and d: "d_\ p l" and lp: "l > 0" - shows "iszlfm (a_\ p l) (a #bs) \ d_\ (a_\ p l) 1 \ (Ifm (real (l * x) #bs) (a_\ p l) = Ifm ((real x)#bs) p)" + shows "iszlfm (a_\ p l) (a #bs) \ d_\ (a_\ p l) 1 \ (Ifm (real_of_int (l * x) #bs) (a_\ p l) = Ifm ((real_of_int x)#bs) p)" using linp d proof (induct p rule: iszlfm.induct) case (5 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ @@ -2438,13 +2416,13 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0::real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)" + hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < (0::real)) = + (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < 0)" by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: algebra_simps) - also have "\ = (real c * real x + Inum (real x # bs) e < 0)" - using mult_less_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp + also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) < (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) + also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0)" + using mult_less_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (6 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2456,13 +2434,13 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0::real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" + hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = + (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: algebra_simps) - also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" - using mult_le_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp + also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \ (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) + also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0)" + using mult_le_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (7 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2474,13 +2452,13 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0::real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)" + hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > (0::real)) = + (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > 0)" by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: algebra_simps) - also have "\ = (real c * real x + Inum (real x # bs) e > 0)" - using zero_less_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp + also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) > (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) + also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e > 0)" + using zero_less_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (8 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2492,13 +2470,13 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0::real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" + hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = + (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: algebra_simps) - also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp + also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \ (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) + also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0)" + using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (3 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2510,13 +2488,13 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0::real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)" + hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (0::real)) = + (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = 0)" by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: algebra_simps) - also have "\ = (real c * real x + Inum (real x # bs) e = 0)" - using mult_eq_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp + also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) = (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) + also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = 0)" + using mult_eq_0_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (4 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2528,13 +2506,13 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(real l * real x + real (l div c) * Inum (real x # bs) e \ (0::real)) = - (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \ 0)" + hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = + (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" by simp - also have "\ = (real (l div c) * (real c * real x + Inum (real x # bs) e) \ (real (l div c)) * 0)" by (simp add: algebra_simps) - also have "\ = (real c * real x + Inum (real x # bs) e \ 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp + also have "\ = (real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e) \ (real_of_int (l div c)) * 0)" by (simp add: algebra_simps) + also have "\ = (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0)" + using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e"] ldcp by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] be isint_Mul[OF ei] by simp next case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2546,12 +2524,12 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(\ (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\ (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp - also have "\ = (\ (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: algebra_simps) - also fix k have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp - also have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp + hence "(\ (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\ (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)" by simp + also have "\ = (\ (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps) + also fix k have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)" + using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp + also have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp next case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) @@ -2563,16 +2541,16 @@ have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp - hence "(\ (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\ (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp - also have "\ = (\ (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: algebra_simps) - also fix k have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" - using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp - also have "\ = (\ (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp - finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp -qed (simp_all add: numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult) + hence "(\ (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\ (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)" by simp + also have "\ = (\ (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps) + also fix k have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)" + using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp + also have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp +qed (simp_all add: numbound0_I[where bs="bs" and b="real_of_int (l * x)" and b'="real_of_int x"] isint_Mul del: of_int_mult) lemma a_\_ex: assumes linp: "iszlfm p (a#bs)" and d: "d_\ p l" and lp: "l>0" - shows "(\ x. l dvd x \ Ifm (real x #bs) (a_\ p l)) = (\ (x::int). Ifm (real x#bs) p)" + shows "(\ x. l dvd x \ Ifm (real_of_int x #bs) (a_\ p l)) = (\ (x::int). Ifm (real_of_int x#bs) p)" (is "(\ x. l dvd x \ ?P x) = (\ x. ?P' x)") proof- have "(\ x. l dvd x \ ?P x) = (\ (x::int). ?P (l*x))" @@ -2586,148 +2564,146 @@ and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" - and nob: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real x = b + real j)" - and p: "Ifm (real x#bs) p" (is "?P x") + and nob: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real_of_int x = b + real_of_int j)" + and p: "Ifm (real_of_int x#bs) p" (is "?P x") shows "?P (x - d)" using lp u d dp nob p proof(induct p rule: iszlfm.induct) case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all - with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] 5 - show ?case by (simp del: real_of_int_minus) + with dp p c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] 5 + show ?case by (simp del: of_int_minus) next case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all - with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] 6 - show ?case by (simp del: real_of_int_minus) + with dp p c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] 6 + show ?case by (simp del: of_int_minus) next - case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1" + case (7 c e) hence p: "Ifm (real_of_int x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp_all - let ?e = "Inum (real x # bs) e" - from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"] - numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"] + let ?e = "Inum (real_of_int x # bs) e" + from ie1 have ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"] + numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"] by (simp add: isint_iff) - {assume "real (x-d) +?e > 0" hence ?case using c1 - numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] - by (simp del: real_of_int_minus)} + {assume "real_of_int (x-d) +?e > 0" hence ?case using c1 + numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] + by (simp del: of_int_minus)} moreover - {assume H: "\ real (x-d) + ?e > 0" + {assume H: "\ real_of_int (x-d) + ?e > 0" let ?v="Neg e" have vb: "?v \ set (\ (Gt (CN 0 c e)))" by simp - from 7(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] - have nob: "\ (\ j\ {1 ..d}. real x = - ?e + real j)" by auto - from H p have "real x + ?e > 0 \ real x + ?e \ real d" by (simp add: c1) - hence "real (x + floor ?e) > real (0::int) \ real (x + floor ?e) \ real d" + from 7(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] + have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e + real_of_int j)" by auto + from H p have "real_of_int x + ?e > 0 \ real_of_int x + ?e \ real_of_int d" by (simp add: c1) + hence "real_of_int (x + floor ?e) > real_of_int (0::int) \ real_of_int (x + floor ?e) \ real_of_int d" using ie by simp hence "x + floor ?e \ 1 \ x + floor ?e \ d" by simp hence "\ (j::int) \ {1 .. d}. j = x + floor ?e" by simp - hence "\ (j::int) \ {1 .. d}. real x = real (- floor ?e + j)" - by (simp only: real_of_int_inject) (simp add: algebra_simps) - hence "\ (j::int) \ {1 .. d}. real x = - ?e + real j" + hence "\ (j::int) \ {1 .. d}. real_of_int x = real_of_int (- floor ?e + j)" by force + hence "\ (j::int) \ {1 .. d}. real_of_int x = - ?e + real_of_int j" by (simp add: ie[simplified isint_iff]) with nob have ?case by auto} ultimately show ?case by blast next - case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" + case (8 c e) hence p: "Ifm (real_of_int x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" - from ie1 have ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"] + let ?e = "Inum (real_of_int x # bs) e" + from ie1 have ie: "real_of_int (floor ?e) = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"] by (simp add: isint_iff) - {assume "real (x-d) +?e \ 0" hence ?case using c1 - numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] - by (simp del: real_of_int_minus)} + {assume "real_of_int (x-d) +?e \ 0" hence ?case using c1 + numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] + by (simp del: of_int_minus)} moreover - {assume H: "\ real (x-d) + ?e \ 0" + {assume H: "\ real_of_int (x-d) + ?e \ 0" let ?v="Sub (C (- 1)) e" have vb: "?v \ set (\ (Ge (CN 0 c e)))" by simp - from 8(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]] - have nob: "\ (\ j\ {1 ..d}. real x = - ?e - 1 + real j)" by auto - from H p have "real x + ?e \ 0 \ real x + ?e < real d" by (simp add: c1) - hence "real (x + floor ?e) \ real (0::int) \ real (x + floor ?e) < real d" + from 8(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] + have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e - 1 + real_of_int j)" by auto + from H p have "real_of_int x + ?e \ 0 \ real_of_int x + ?e < real_of_int d" by (simp add: c1) + hence "real_of_int (x + floor ?e) \ real_of_int (0::int) \ real_of_int (x + floor ?e) < real_of_int d" using ie by simp hence "x + floor ?e +1 \ 1 \ x + floor ?e + 1 \ d" by simp hence "\ (j::int) \ {1 .. d}. j = x + floor ?e + 1" by simp hence "\ (j::int) \ {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps) - hence "\ (j::int) \ {1 .. d}. real x= real (- floor ?e - 1 + j)" - by (simp only: real_of_int_inject) - hence "\ (j::int) \ {1 .. d}. real x= - ?e - 1 + real j" + hence "\ (j::int) \ {1 .. d}. real_of_int x= real_of_int (- floor ?e - 1 + j)" by blast + hence "\ (j::int) \ {1 .. d}. real_of_int x= - ?e - 1 + real_of_int j" by (simp add: ie[simplified isint_iff]) with nob have ?case by simp } ultimately show ?case by blast next - case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" + case (3 c e) hence p: "Ifm (real_of_int x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" + let ?e = "Inum (real_of_int x # bs) e" let ?v="(Sub (C (- 1)) e)" have vb: "?v \ set (\ (Eq (CN 0 c e)))" by simp - from p have "real x= - ?e" by (simp add: c1) with 3(5) show ?case using dp + from p have "real_of_int x= - ?e" by (simp add: c1) with 3(5) show ?case using dp by simp (erule ballE[where x="1"], - simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"]) + simp_all add:algebra_simps numbound0_I[OF bn,where b="real_of_int x"and b'="a"and bs="bs"]) next - case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" + case (4 c e)hence p: "Ifm (real_of_int x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" + let ?e = "Inum (real_of_int x # bs) e" let ?v="Neg e" have vb: "?v \ set (\ (NEq (CN 0 c e)))" by simp - {assume "real x - real d + Inum ((real (x -d)) # bs) e \ 0" + {assume "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e \ 0" hence ?case by (simp add: c1)} moreover - {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0" - hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp - hence "real x = - Inum (a # bs) e + real d" - by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"]) + {assume H: "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e = 0" + hence "real_of_int x = - Inum ((real_of_int (x -d)) # bs) e + real_of_int d" by simp + hence "real_of_int x = - Inum (a # bs) e + real_of_int d" + by (simp add: numbound0_I[OF bn,where b="real_of_int x - real_of_int d"and b'="a"and bs="bs"]) with 4(5) have ?case using dp by simp} ultimately show ?case by blast next - case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" + case (9 j c e) hence p: "Ifm (real_of_int x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" + let ?e = "Inum (real_of_int x # bs) e" from 9 have "isint e (a #bs)" by simp - hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real x)#bs"] numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] + hence ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real_of_int x)#bs"] numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] by (simp add: isint_iff) from 9 have id: "j dvd d" by simp - from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp + from c1 ie[symmetric] have "?p x = (real_of_int j rdvd real_of_int (x+ floor ?e))" by simp also have "\ = (j dvd x + floor ?e)" - using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp + using int_rdvd_real[where i="j" and x="real_of_int (x+ floor ?e)"] by simp also have "\ = (j dvd x - d + floor ?e)" using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp - also have "\ = (real j rdvd real (x - d + floor ?e))" - using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] + also have "\ = (real_of_int j rdvd real_of_int (x - d + floor ?e))" + using int_rdvd_real[where i="j" and x="real_of_int (x-d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = (real j rdvd real x - real d + ?e)" + also have "\ = (real_of_int j rdvd real_of_int x - real_of_int d + ?e)" using ie by simp finally show ?case - using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp + using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp next - case (10 j c e) hence p: "Ifm (real x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ - let ?e = "Inum (real x # bs) e" + case (10 j c e) hence p: "Ifm (real_of_int x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ + let ?e = "Inum (real_of_int x # bs) e" from 10 have "isint e (a#bs)" by simp - hence ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"] + hence ie: "real_of_int (floor ?e) = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"] by (simp add: isint_iff) from 10 have id: "j dvd d" by simp - from c1 ie[symmetric] have "?p x = (\ real j rdvd real (x+ floor ?e))" by simp + from c1 ie[symmetric] have "?p x = (\ real_of_int j rdvd real_of_int (x+ floor ?e))" by simp also have "\ = (\ j dvd x + floor ?e)" - using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp + using int_rdvd_real[where i="j" and x="real_of_int (x+ floor ?e)"] by simp also have "\ = (\ j dvd x - d + floor ?e)" using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp - also have "\ = (\ real j rdvd real (x - d + floor ?e))" - using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] + also have "\ = (\ real_of_int j rdvd real_of_int (x - d + floor ?e))" + using int_rdvd_real[where i="j" and x="real_of_int (x-d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = (\ real j rdvd real x - real d + ?e)" + also have "\ = (\ real_of_int j rdvd real_of_int x - real_of_int d + ?e)" using ie by simp finally show ?case - using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp -qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"] - simp del: real_of_int_diff) + using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp +qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int (x - d)" and b'="real_of_int x"] + simp del: of_int_diff) lemma \': assumes lp: "iszlfm p (a #bs)" and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" - shows "\ x. \(\(j::int) \ {1 .. d}. \ b\ set(\ p). Ifm ((Inum (a#bs) b + real j) #bs) p) \ Ifm (real x#bs) p \ Ifm (real (x - d)#bs) p" (is "\ x. ?b \ ?P x \ ?P (x - d)") + shows "\ x. \(\(j::int) \ {1 .. d}. \ b\ set(\ p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p) \ Ifm (real_of_int x#bs) p \ Ifm (real_of_int (x - d)#bs) p" (is "\ x. ?b \ ?P x \ ?P (x - d)") proof(clarify) fix x assume nb:"?b" and px: "?P x" - hence nb2: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real x = b + real j)" + hence nb2: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real_of_int x = b + real_of_int j)" by auto from \[OF lp u d dp nb2 px] show "?P (x -d )" . qed @@ -2764,22 +2740,22 @@ and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" - shows "(\ (x::int). Ifm (real x #bs) p) = (\ j\ {1.. d}. Ifm (real j #bs) (minusinf p) \ (\ b \ set (\ p). Ifm ((Inum (a#bs) b + real j) #bs) p))" - (is "(\ (x::int). ?P (real x)) = (\ j\ ?D. ?M j \ (\ b\ ?B. ?P (?I b + real j)))") + shows "(\ (x::int). Ifm (real_of_int x #bs) p) = (\ j\ {1.. d}. Ifm (real_of_int j #bs) (minusinf p) \ (\ b \ set (\ p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p))" + (is "(\ (x::int). ?P (real_of_int x)) = (\ j\ ?D. ?M j \ (\ b\ ?B. ?P (?I b + real_of_int j)))") proof- from minusinf_inf[OF lp] - have th: "\(z::int). \x(z::int). \x_int[OF lp] isint_iff[where bs="a # bs"] have B: "\ b\ ?B. real (floor (?I b)) = ?I b" by simp + from \_int[OF lp] isint_iff[where bs="a # bs"] have B: "\ b\ ?B. real_of_int (floor (?I b)) = ?I b" by simp from B[rule_format] - have "(\j\?D. \b\ ?B. ?P (?I b + real j)) = (\j\?D. \b\ ?B. ?P (real (floor (?I b)) + real j))" + have "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\j\?D. \b\ ?B. ?P (real_of_int (floor (?I b)) + real_of_int j))" by simp - also have "\ = (\j\?D. \b\ ?B. ?P (real (floor (?I b) + j)))" by simp - also have"\ = (\ j \ ?D. \ b \ ?B'. ?P (real (b + j)))" by blast + also have "\ = (\j\?D. \b\ ?B. ?P (real_of_int (floor (?I b) + j)))" by simp + also have"\ = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" by blast finally have BB': - "(\j\?D. \b\ ?B. ?P (?I b + real j)) = (\ j \ ?D. \ b \ ?B'. ?P (real (b + j)))" + "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" by blast - hence th2: "\ x. \ (\ j \ ?D. \ b \ ?B'. ?P (real (b + j))) \ ?P (real x) \ ?P (real (x - d))" using \'[OF lp u d dp] by blast + hence th2: "\ x. \ (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j))) \ ?P (real_of_int x) \ ?P (real_of_int (x - d))" using \'[OF lp u d dp] by blast from minusinf_repeats[OF d lp] have th3: "\ x k. ?M x = ?M (x-k*d)" by simp from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast @@ -2840,38 +2816,38 @@ "\ p k t \ And (Dvd k t) (\_\ p (t,k))" lemma \_\: - assumes linp: "iszlfm p (real (x::int)#bs)" - and kpos: "real k > 0" + assumes linp: "iszlfm p (real_of_int (x::int)#bs)" + and kpos: "real_of_int k > 0" and tnb: "numbound0 t" - and tint: "isint t (real x#bs)" + and tint: "isint t (real_of_int x#bs)" and kdt: "k dvd floor (Inum (b'#bs) t)" - shows "Ifm (real x#bs) (\_\ p (t,k)) = - (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" - (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)") + shows "Ifm (real_of_int x#bs) (\_\ p (t,k)) = + (Ifm ((real_of_int ((floor (Inum (b'#bs) t)) div k))#bs) p)" + (is "?I (real_of_int x) (?s p) = (?I (real_of_int ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)") using linp kpos tnb proof(induct p rule: \_\.induct) case (3 c e) from 3 have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from kpos have knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" + from kpos have knz': "real_of_int k \ 0" by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp - from assms * have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)" + from assms * have "?I (real_of_int x) (?s (Eq (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k = 0)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', - where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -2879,24 +2855,24 @@ case (4 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from kpos have knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" + from kpos have knz': "real_of_int k \ 0" by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (NEq (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', - where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -2904,23 +2880,23 @@ case (5 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)" + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (Lt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k < 0)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, - where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -2928,23 +2904,23 @@ case (6 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (Le (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, - where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -2952,23 +2928,23 @@ case (7 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)" + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (Gt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k > 0)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, - where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -2976,23 +2952,23 @@ case (8 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \ 0)" + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (Ge (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, - where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + where a="real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e" and b="0", symmetric] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -3000,23 +2976,23 @@ case (9 i c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)" + from kpos have knz: "k\0" by simp hence knz': "real_of_int k \ 0" by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (Dvd i (CN 0 c e))) = (real_of_int i * real_of_int k rdvd (real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k)" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] - real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast @@ -3024,28 +3000,28 @@ case (10 i c e) then have cp: "c > 0" and nb: "numbound0 e" by auto { assume kdc: "k dvd c" - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } moreover { assume *: "\ k dvd c" - from kpos have knz: "k\0" by simp hence knz': "real k \ 0" by simp - from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp - from assms * have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\ (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))" + from kpos have knz: "k\0" by simp hence knz': "real_of_int k \ 0" by simp + from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp + from assms * have "?I (real_of_int x) (?s (NDvd i (CN 0 c e))) = (\ (real_of_int i * real_of_int k rdvd (real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k))" using real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti algebra_simps) also have "\ = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF kdt] - numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] + numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } ultimately show ?case by blast -qed (simp_all add: bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"] - numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"]) +qed (simp_all add: bound0_I[where bs="bs" and b="real_of_int ((floor (?N b' t)) div k)" and b'="real_of_int x"] + numbound0_I[where bs="bs" and b="real_of_int ((floor (?N b' t)) div k)" and b'="real_of_int x"]) lemma \_\_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t" @@ -3054,153 +3030,153 @@ by (induct p rule: iszlfm.induct, auto simp add: nb) lemma \_l: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "\ (b,k) \ set (\ p). k >0 \ numbound0 b \ isint b (real i#bs)" + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" + shows "\ (b,k) \ set (\ p). k >0 \ numbound0 b \ isint b (real_of_int i#bs)" using lp by (induct p rule: \.induct, auto simp add: isint_sub isint_neg) lemma \_\_l: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "\ (b,k) \ set (\_\ p). k >0 \ numbound0 b \ isint b (real i#bs)" -using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"] + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" + shows "\ (b,k) \ set (\_\ p). k >0 \ numbound0 b \ isint b (real_of_int i#bs)" +using lp isint_add [OF isint_c[where j="- 1"],where bs="real_of_int i#bs"] by (induct p rule: \_\.induct, auto) -lemma \: assumes lp: "iszlfm p (real (i::int) #bs)" - and pi: "Ifm (real i#bs) p" +lemma \: assumes lp: "iszlfm p (real_of_int (i::int) #bs)" + and pi: "Ifm (real_of_int i#bs) p" and d: "d_\ p d" and dp: "d > 0" - and nob: "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. real (c*i) \ Inum (real i#bs) e + real j" + and nob: "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. real_of_int (c*i) \ Inum (real_of_int i#bs) e + real_of_int j" (is "\(e,c) \ set (\ p). \ j\ {1 .. c*d}. _ \ ?N i e + _") - shows "Ifm (real(i - d)#bs) p" + shows "Ifm (real_of_int(i - d)#bs) p" using lp pi d nob proof(induct p rule: iszlfm.induct) - case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and pi: "real (c*i) = - 1 - ?N i e + real (1::int)" and nob: "\ j\ {1 .. c*d}. real (c*i) \ -1 - ?N i e + real j" + case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" + and pi: "real_of_int (c*i) = - 1 - ?N i e + real_of_int (1::int)" and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ -1 - ?N i e + real_of_int j" by simp+ from mult_strict_left_mono[OF dp cp] have one:"1 \ {1 .. c*d}" by auto from nob[rule_format, where j="1", OF one] pi show ?case by simp next case (4 c e) - hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and nob: "\ j\ {1 .. c*d}. real (c*i) \ - ?N i e + real j" + hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" + and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - ?N i e + real_of_int j" by simp+ - {assume "real (c*i) \ - ?N i e + real (c*d)" - with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"] + {assume "real_of_int (c*i) \ - ?N i e + real_of_int (c*d)" + with numbound0_I[OF nb, where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] have ?case by (simp add: algebra_simps)} moreover - {assume pi: "real (c*i) = - ?N i e + real (c*d)" + {assume pi: "real_of_int (c*i) = - ?N i e + real_of_int (c*d)" from mult_strict_left_mono[OF dp cp] have d: "(c*d) \ {1 .. c*d}" by simp from nob[rule_format, where j="c*d", OF d] pi have ?case by simp } ultimately show ?case by blast next case (5 c e) hence cp: "c > 0" by simp - from 5 mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] - real_of_int_mult] + from 5 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] + of_int_mult] show ?case using 5 dp - by (simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] + apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] algebra_simps del: mult_pos_pos) + by (metis add.right_neutral of_int_0_less_iff of_int_mult pos_add_strict) next case (6 c e) hence cp: "c > 0" by simp - from 6 mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] - real_of_int_mult] + from 6 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] + of_int_mult] show ?case using 6 dp - by (simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] + apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] algebra_simps del: mult_pos_pos) + using order_trans by fastforce next - case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and nob: "\ j\ {1 .. c*d}. real (c*i) \ - ?N i e + real j" - and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0" + case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" + and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - ?N i e + real_of_int j" + and pi: "real_of_int (c*i) + ?N i e > 0" and cp': "real_of_int c >0" by simp+ let ?fe = "floor (?N i e)" - from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: algebra_simps) - from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp - hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric]) - have "real (c*i) + ?N i e > real (c*d) \ real (c*i) + ?N i e \ real (c*d)" by auto + from pi cp have th:"(real_of_int i +?N i e / real_of_int c)*real_of_int c > 0" by (simp add: algebra_simps) + from pi ei[simplified isint_iff] have "real_of_int (c*i + ?fe) > real_of_int (0::int)" by simp + hence pi': "c*i + ?fe > 0" by (simp only: of_int_less_iff[symmetric]) + have "real_of_int (c*i) + ?N i e > real_of_int (c*d) \ real_of_int (c*i) + ?N i e \ real_of_int (c*d)" by auto moreover - {assume "real (c*i) + ?N i e > real (c*d)" hence ?case + {assume "real_of_int (c*i) + ?N i e > real_of_int (c*d)" hence ?case by (simp add: algebra_simps - numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} + numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} moreover - {assume H:"real (c*i) + ?N i e \ real (c*d)" - with ei[simplified isint_iff] have "real (c*i + ?fe) \ real (c*d)" by simp - hence pid: "c*i + ?fe \ c*d" by (simp only: real_of_int_le_iff) + {assume H:"real_of_int (c*i) + ?N i e \ real_of_int (c*d)" + with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \ real_of_int (c*d)" by simp + hence pid: "c*i + ?fe \ c*d" by (simp only: of_int_le_iff) with pi' have "\ j1\ {1 .. c*d}. c*i + ?fe = j1" by auto - hence "\ j1\ {1 .. c*d}. real (c*i) = - ?N i e + real j1" - by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff]) - (simp add: algebra_simps) + hence "\ j1\ {1 .. c*d}. real_of_int (c*i) = - ?N i e + real_of_int j1" + unfolding Bex_def using ei[simplified isint_iff] by fastforce with nob have ?case by blast } ultimately show ?case by blast next - case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" - and nob: "\ j\ {1 .. c*d}. real (c*i) \ - 1 - ?N i e + real j" - and pi: "real (c*i) + ?N i e \ 0" and cp': "real c >0" + case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" + and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - 1 - ?N i e + real_of_int j" + and pi: "real_of_int (c*i) + ?N i e \ 0" and cp': "real_of_int c >0" by simp+ let ?fe = "floor (?N i e)" - from pi cp have th:"(real i +?N i e / real c)*real c \ 0" by (simp add: algebra_simps) - from pi ei[simplified isint_iff] have "real (c*i + ?fe) \ real (0::int)" by simp - hence pi': "c*i + 1 + ?fe \ 1" by (simp only: real_of_int_le_iff[symmetric]) - have "real (c*i) + ?N i e \ real (c*d) \ real (c*i) + ?N i e < real (c*d)" by auto + from pi cp have th:"(real_of_int i +?N i e / real_of_int c)*real_of_int c \ 0" by (simp add: algebra_simps) + from pi ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \ real_of_int (0::int)" by simp + hence pi': "c*i + 1 + ?fe \ 1" by (simp only: of_int_le_iff[symmetric]) + have "real_of_int (c*i) + ?N i e \ real_of_int (c*d) \ real_of_int (c*i) + ?N i e < real_of_int (c*d)" by auto moreover - {assume "real (c*i) + ?N i e \ real (c*d)" hence ?case + {assume "real_of_int (c*i) + ?N i e \ real_of_int (c*d)" hence ?case by (simp add: algebra_simps - numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} + numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} moreover - {assume H:"real (c*i) + ?N i e < real (c*d)" - with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp - hence pid: "c*i + 1 + ?fe \ c*d" by (simp only: real_of_int_le_iff) + {assume H:"real_of_int (c*i) + ?N i e < real_of_int (c*d)" + with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) < real_of_int (c*d)" by simp + hence pid: "c*i + 1 + ?fe \ c*d" by (simp only: of_int_le_iff) with pi' have "\ j1\ {1 .. c*d}. c*i + 1+ ?fe = j1" by auto - hence "\ j1\ {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1" - by (simp only: real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] real_of_one) - (simp add: algebra_simps) - hence "\ j1\ {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1" + hence "\ j1\ {1 .. c*d}. real_of_int (c*i) + 1= - ?N i e + real_of_int j1" + unfolding Bex_def using ei[simplified isint_iff] by fastforce + hence "\ j1\ {1 .. c*d}. real_of_int (c*i) = (- ?N i e + real_of_int j1) - 1" by (simp only: algebra_simps) - hence "\ j1\ {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1" + hence "\ j1\ {1 .. c*d}. real_of_int (c*i) = - 1 - ?N i e + real_of_int j1" by (simp add: algebra_simps) with nob have ?case by blast } ultimately show ?case by blast next - case (9 j c e) hence p: "real j rdvd real (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ - let ?e = "Inum (real i # bs) e" - from 9 have "isint e (real i #bs)" by simp - hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"] + case (9 j c e) hence p: "real_of_int j rdvd real_of_int (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ + let ?e = "Inum (real_of_int i # bs) e" + from 9 have "isint e (real_of_int i #bs)" by simp + hence ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"] by (simp add: isint_iff) from 9 have id: "j dvd d" by simp - from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp + from ie[symmetric] have "?p i = (real_of_int j rdvd real_of_int (c*i+ floor ?e))" by simp also have "\ = (j dvd c*i + floor ?e)" using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp also have "\ = (j dvd c*i - c*d + floor ?e)" using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp - also have "\ = (real j rdvd real (c*i - c*d + floor ?e))" + also have "\ = (real_of_int j rdvd real_of_int (c*i - c*d + floor ?e))" using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = (real j rdvd real (c*(i - d)) + ?e)" + also have "\ = (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" using ie by (simp add:algebra_simps) finally show ?case - using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p + using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p by (simp add: algebra_simps) next case (10 j c e) - hence p: "\ (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" + hence p: "\ (real_of_int j rdvd real_of_int (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ - let ?e = "Inum (real i # bs) e" - from 10 have "isint e (real i #bs)" by simp - hence ie: "real (floor ?e) = ?e" - using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"] + let ?e = "Inum (real_of_int i # bs) e" + from 10 have "isint e (real_of_int i #bs)" by simp + hence ie: "real_of_int (floor ?e) = ?e" + using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"] by (simp add: isint_iff) from 10 have id: "j dvd d" by simp - from ie[symmetric] have "?p i = (\ (real j rdvd real (c*i+ floor ?e)))" by simp + from ie[symmetric] have "?p i = (\ (real_of_int j rdvd real_of_int (c*i+ floor ?e)))" by simp also have "\ = Not (j dvd c*i + floor ?e)" using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp also have "\ = Not (j dvd c*i - c*d + floor ?e)" using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp - also have "\ = Not (real j rdvd real (c*i - c*d + floor ?e))" + also have "\ = Not (real_of_int j rdvd real_of_int (c*i - c*d + floor ?e))" using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = Not (real j rdvd real (c*(i - d)) + ?e)" + also have "\ = Not (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" using ie by (simp add:algebra_simps) finally show ?case - using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p + using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p by (simp add: algebra_simps) -qed (auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"]) +qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]) lemma \_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t" shows "bound0 (\ p k t)" @@ -3209,37 +3185,37 @@ lemma \': assumes lp: "iszlfm p (a #bs)" and d: "d_\ p d" and dp: "d > 0" - shows "\ x. \(\ (e,c) \ set(\ p). \(j::int) \ {1 .. c*d}. Ifm (a #bs) (\ p c (Add e (C j)))) \ Ifm (real x#bs) p \ Ifm (real (x - d)#bs) p" (is "\ x. ?b x \ ?P x \ ?P (x - d)") + shows "\ x. \(\ (e,c) \ set(\ p). \(j::int) \ {1 .. c*d}. Ifm (a #bs) (\ p c (Add e (C j)))) \ Ifm (real_of_int x#bs) p \ Ifm (real_of_int (x - d)#bs) p" (is "\ x. ?b x \ ?P x \ ?P (x - d)") proof(clarify) fix x assume nob1:"?b x" and px: "?P x" - from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)". - have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real (c * x) \ Inum (real x # bs) e + real j" + from iszlfm_gen[OF lp, rule_format, where y="real_of_int x"] have lp': "iszlfm p (real_of_int x#bs)". + have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real_of_int (c * x) \ Inum (real_of_int x # bs) e + real_of_int j" proof(clarify) fix e c j assume ecR: "(e,c) \ set (\ p)" and jD: "j\ {1 .. c*d}" - and cx: "real (c*x) = Inum (real x#bs) e + real j" - let ?e = "Inum (real x#bs) e" + and cx: "real_of_int (c*x) = Inum (real_of_int x#bs) e + real_of_int j" + let ?e = "Inum (real_of_int x#bs) e" let ?fe = "floor ?e" - from \_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e" + from \_l[OF lp'] ecR have ei:"isint e (real_of_int x#bs)" and cp:"c>0" and nb:"numbound0 e" by auto from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" . - from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp - hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject) + from cx ei[simplified isint_iff] have "real_of_int (c*x) = real_of_int (?fe + j)" by simp + hence cx: "c*x = ?fe + j" by (simp only: of_int_eq_iff) hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp) - hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff) - hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff]) + hence "real_of_int c rdvd real_of_int (?fe + j)" by (simp only: int_rdvd_iff) + hence rcdej: "real_of_int c rdvd ?e + real_of_int j" by (simp add: ei[simplified isint_iff]) from cx have "(c*x) div c = (?fe + j) div c" by simp with cp have "x = (?fe + j) div c" by simp with px have th: "?P ((?fe + j) div c)" by auto - from cp have cp': "real c > 0" by simp - from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp + from cp have cp': "real_of_int c > 0" by simp + from cdej have cdej': "c dvd floor (Inum (real_of_int x#bs) (Add e (C j)))" by simp from nb have nb': "numbound0 (Add e (C j))" by simp - have ji: "isint (C j) (real x#bs)" by (simp add: isint_def) - from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" . - from th \_\[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric] - have "Ifm (real x#bs) (\_\ p (Add e (C j), c))" by simp - with rcdej have th: "Ifm (real x#bs) (\ p c (Add e (C j)))" by (simp add: \_def) - from th bound0_I[OF \_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"] + have ji: "isint (C j) (real_of_int x#bs)" by (simp add: isint_def) + from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real_of_int x#bs)" . + from th \_\[where b'="real_of_int x", OF lp' cp' nb' ei' cdej',symmetric] + have "Ifm (real_of_int x#bs) (\_\ p (Add e (C j), c))" by simp + with rcdej have th: "Ifm (real_of_int x#bs) (\ p c (Add e (C j)))" by (simp add: \_def) + from th bound0_I[OF \_nb[OF lp nb', where k="c"],where bs="bs" and b="real_of_int x" and b'="a"] have "Ifm (a#bs) (\ p c (Add e (C j)))" by blast with ecR jD nob1 show "False" by blast qed @@ -3248,8 +3224,8 @@ lemma rl_thm: - assumes lp: "iszlfm p (real (i::int)#bs)" - shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ set (\ p). \ j\ {1 .. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" + shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ (e,c) \ set (\ p). \ j\ {1 .. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" (is "(\(x::int). ?P x) = ((\ j\ {1.. \ p}. ?MP j)\(\ (e,c) \ ?R. \ j\ _. ?SP c e j))" is "?lhs = (?MD \ ?RD)" is "?lhs = ?rhs") proof- @@ -3259,21 +3235,21 @@ from H minusinf_ex[OF lp th] have ?thesis by blast} moreover { fix e c j assume exR:"(e,c) \ ?R" and jD:"j\ {1 .. c*?d}" and spx:"?SP c e j" - from exR \_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0" + from exR \_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real_of_int i#bs)" and cp: "c > 0" by auto - have "isint (C j) (real i#bs)" by (simp add: isint_iff) - with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]] - have eji:"isint (Add e (C j)) (real i#bs)" by simp + have "isint (C j) (real_of_int i#bs)" by (simp add: isint_iff) + with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real_of_int i"]] + have eji:"isint (Add e (C j)) (real_of_int i#bs)" by simp from nb have nb': "numbound0 (Add e (C j))" by simp - from spx bound0_I[OF \_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"] - have spx': "Ifm (real i # bs) (\ p c (Add e (C j)))" by blast - from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" - and sr:"Ifm (real i#bs) (\_\ p (Add e (C j),c))" by (simp add: \_def)+ + from spx bound0_I[OF \_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real_of_int i"] + have spx': "Ifm (real_of_int i # bs) (\ p c (Add e (C j)))" by blast + from spx' have rcdej:"real_of_int c rdvd (Inum (real_of_int i#bs) (Add e (C j)))" + and sr:"Ifm (real_of_int i#bs) (\_\ p (Add e (C j),c))" by (simp add: \_def)+ from rcdej eji[simplified isint_iff] - have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp - hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff) - from cp have cp': "real c > 0" by simp - from \_\[OF lp cp' nb' eji cdej] spx' have "?P (\Inum (real i # bs) (Add e (C j))\ div c)" + have "real_of_int c rdvd real_of_int (floor (Inum (real_of_int i#bs) (Add e (C j))))" by simp + hence cdej:"c dvd floor (Inum (real_of_int i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff) + from cp have cp': "real_of_int c > 0" by simp + from \_\[OF lp cp' nb' eji cdej] spx' have "?P (\Inum (real_of_int i # bs) (Add e (C j))\ div c)" by (simp add: \_def) hence ?lhs by blast with exR jD spx have ?thesis by blast} @@ -3440,10 +3416,10 @@ let ?l = "floor (?N s') + j" from H have "?I (?p (p',n',s') j) \ - (((?N ?nxs \ real ?l) \ (?N ?nxs < real (?l + 1))) \ (?N a = ?N ?nxs ))" + (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" by (simp add: fp_def np algebra_simps) also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" - using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp + using floor_unique_iff[where x="?N ?nxs" and a="?l"] by simp moreover have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" @@ -3466,10 +3442,10 @@ let ?l = "floor (?N s') + j" from H have "?I (?p (p',n',s') j) \ - (((?N ?nxs \ real ?l) \ (?N ?nxs < real (?l + 1))) \ (?N a = ?N ?nxs ))" + (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" by (simp add: np fp_def algebra_simps) also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" - using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp + using floor_unique_iff[where x="?N ?nxs" and a="?l"] by simp moreover have "\ \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp ultimately have "?I (?p (p',n',s') j) \ (?N (Floor a) = ?N ((Add (Floor s') (C j))))" @@ -3483,10 +3459,11 @@ qed (auto simp add: Let_def split_def algebra_simps) lemma real_in_int_intervals: - assumes xb: "real m \ x \ x < real ((n::int) + 1)" - shows "\ j\ {m.. n}. real j \ x \ x < real (j+1)" (is "\ j\ ?N. ?P j") + assumes xb: "real_of_int m \ x \ x < real_of_int ((n::int) + 1)" + shows "\ j\ {m.. n}. real_of_int j \ x \ x < real_of_int (j+1)" (is "\ j\ ?N. ?P j") by (rule bexI[where P="?P" and x="floor x" and A="?N"]) -(auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]]) +(auto simp add: floor_less_iff[where x="x" and z="n+1", simplified] + xb[simplified] floor_mono[where x="real_of_int m" and y="x", OF conjunct1[OF xb], simplified floor_of_int[where z="m"]]) lemma rsplit0_complete: assumes xp:"0 \ x" and x1:"x < 1" @@ -3571,20 +3548,20 @@ moreover { assume np: "n > 0" - from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \ ?N s" by simp - also from mult_left_mono[OF xp] np have "?N s \ real n * x + ?N s" by simp - finally have "?N (Floor s) \ real n * x + ?N s" . + from of_int_floor_le[of "?N s"] have "?N (Floor s) \ ?N s" by simp + also from mult_left_mono[OF xp] np have "?N s \ real_of_int n * x + ?N s" by simp + finally have "?N (Floor s) \ real_of_int n * x + ?N s" . moreover - {from x1 np have "real n *x + ?N s < real n + ?N s" by simp + {from x1 np have "real_of_int n *x + ?N s < real_of_int n + ?N s" by simp also from real_of_int_floor_add_one_gt[where r="?N s"] - have "\ < real n + ?N (Floor s) + 1" by simp - finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp} - ultimately have "?N (Floor s) \ real n *x + ?N s\ real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp - hence th: "0 \ real n *x + ?N s - ?N (Floor s) \ real n *x + ?N s - ?N (Floor s) < real (n+1)" by simp - from real_in_int_intervals th have "\ j\ {0 .. n}. real j \ real n *x + ?N s - ?N (Floor s)\ real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp + have "\ < real_of_int n + ?N (Floor s) + 1" by simp + finally have "real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp} + ultimately have "?N (Floor s) \ real_of_int n *x + ?N s\ real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp + hence th: "0 \ real_of_int n *x + ?N s - ?N (Floor s) \ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (n+1)" by simp + from real_in_int_intervals th have "\ j\ {0 .. n}. real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s)\ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp - hence "\ j\ {0 .. n}. 0 \ real n *x + ?N s - ?N (Floor s) - real j \ real n *x + ?N s - ?N (Floor s) - real (j+1) < 0" - by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) + hence "\ j\ {0 .. n}. 0 \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0" + by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) hence "\ j\ {0.. n}. ?I (?p (p,n,s) j)" using pns by (simp add: fp_def np algebra_simps) then obtain "j" where j_def: "j\ {0 .. n} \ ?I (?p (p,n,s) j)" by blast @@ -3596,23 +3573,23 @@ } moreover { assume nn: "n < 0" hence np: "-n >0" by simp - from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp - moreover from mult_left_mono_neg[OF xp] nn have "?N s \ real n * x + ?N s" by simp - ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith + from of_int_floor_le[of "?N s"] have "?N (Floor s) + 1 > ?N s" by simp + moreover from mult_left_mono_neg[OF xp] nn have "?N s \ real_of_int n * x + ?N s" by simp + ultimately have "?N (Floor s) + 1 > real_of_int n * x + ?N s" by arith moreover - {from x1 nn have "real n *x + ?N s \ real n + ?N s" by simp - moreover from real_of_int_floor_le[where r="?N s"] have "real n + ?N s \ real n + ?N (Floor s)" by simp - ultimately have "real n *x + ?N s \ ?N (Floor s) + real n" + {from x1 nn have "real_of_int n *x + ?N s \ real_of_int n + ?N s" by simp + moreover from of_int_floor_le[of "?N s"] have "real_of_int n + ?N s \ real_of_int n + ?N (Floor s)" by simp + ultimately have "real_of_int n *x + ?N s \ ?N (Floor s) + real_of_int n" by (simp only: algebra_simps)} - ultimately have "?N (Floor s) + real n \ real n *x + ?N s\ real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp - hence th: "real n \ real n *x + ?N s - ?N (Floor s) \ real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp + ultimately have "?N (Floor s) + real_of_int n \ real_of_int n *x + ?N s\ real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (1::int)" by simp + hence th: "real_of_int n \ real_of_int n *x + ?N s - ?N (Floor s) \ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (1::int)" by simp have th1: "\ (a::real). (- a > 0) = (a < 0)" by auto have th2: "\ (a::real). (0 \ - a) = (a \ 0)" by auto - from real_in_int_intervals th have "\ j\ {n .. 0}. real j \ real n *x + ?N s - ?N (Floor s)\ real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp + from real_in_int_intervals th have "\ j\ {n .. 0}. real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s)\ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp - hence "\ j\ {n .. 0}. 0 \ real n *x + ?N s - ?N (Floor s) - real j \ real n *x + ?N s - ?N (Floor s) - real (j+1) < 0" - by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"]) - hence "\ j\ {n .. 0}. 0 \ - (real n *x + ?N s - ?N (Floor s) - real j) \ - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format]) + hence "\ j\ {n .. 0}. 0 \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0" + by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) + hence "\ j\ {n .. 0}. 0 \ - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j) \ - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format]) hence "\ j\ {n.. 0}. ?I (?p (p,n,s) j)" using pns by (simp add: fp_def nn algebra_simps del: diff_less_0_iff_less diff_le_0_iff_le) @@ -3773,37 +3750,37 @@ lemma small_le: assumes u0:"0 \ u" and u1: "u < 1" - shows "(-u \ real (n::int)) = (0 \ n)" + shows "(-u \ real_of_int (n::int)) = (0 \ n)" using u0 u1 by auto lemma small_lt: assumes u0:"0 \ u" and u1: "u < 1" - shows "(real (n::int) < real (m::int) - u) = (n < m)" + shows "(real_of_int (n::int) < real_of_int (m::int) - u) = (n < m)" using u0 u1 by auto lemma rdvd01_cs: - assumes up: "u \ 0" and u1: "u<1" and np: "real n > 0" - shows "(real (i::int) rdvd real (n::int) * u - s) = (\ j\ {0 .. n - 1}. real n * u = s - real (floor s) + real j \ real i rdvd real (j - floor s))" (is "?lhs = ?rhs") + assumes up: "u \ 0" and u1: "u<1" and np: "real_of_int n > 0" + shows "(real_of_int (i::int) rdvd real_of_int (n::int) * u - s) = (\ j\ {0 .. n - 1}. real_of_int n * u = s - real_of_int (floor s) + real_of_int j \ real_of_int i rdvd real_of_int (j - floor s))" (is "?lhs = ?rhs") proof- - let ?ss = "s - real (floor s)" + let ?ss = "s - real_of_int (floor s)" from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]] - real_of_int_floor_le[where r="s"] have ss0:"?ss \ 0" and ss1:"?ss < 1" + of_int_floor_le have ss0:"?ss \ 0" and ss1:"?ss < 1" by (auto simp add: myle[of _ "s", symmetric] myless[of "?ss"]) - from np have n0: "real n \ 0" by simp + from np have n0: "real_of_int n \ 0" by simp from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] - have nu0:"real n * u - s \ -s" and nun:"real n * u -s < real n - s" by auto - from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] - have "real i rdvd real n * u - s = - (i dvd floor (real n * u -s) \ (real (floor (real n * u - s)) = real n * u - s ))" + have nu0:"real_of_int n * u - s \ -s" and nun:"real_of_int n * u -s < real_of_int n - s" by auto + from int_rdvd_real[where i="i" and x="real_of_int (n::int) * u - s"] + have "real_of_int i rdvd real_of_int n * u - s = + (i dvd floor (real_of_int n * u -s) \ (real_of_int (floor (real_of_int n * u - s)) = real_of_int n * u - s ))" (is "_ = (?DE)" is "_ = (?D \ ?E)") by simp - also have "\ = (?DE \ real(floor (real n * u - s) + floor s)\ -?ss - \ real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \real ?a \ _ \ real ?a < _)") + also have "\ = (?DE \ real_of_int(floor (real_of_int n * u - s) + floor s)\ -?ss + \ real_of_int(floor (real_of_int n * u - s) + floor s)< real_of_int n - ?ss)" (is "_=(?DE \real_of_int ?a \ _ \ real_of_int ?a < _)") using nu0 nun by auto also have "\ = (?DE \ ?a \ 0 \ ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1]) also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. ?a = j))" by simp - also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. real (\real n * u - s\) = real j - real \s\ ))" - by (simp only: algebra_simps real_of_int_diff[symmetric] real_of_int_inject) - also have "\ = ((\ j\ {0 .. (n - 1)}. real n * u - s = real j - real \s\ \ real i rdvd real n * u - s))" using int_rdvd_iff[where i="i" and t="\real n * u - s\"] + also have "\ = (?DE \ (\ j\ {0 .. (n - 1)}. real_of_int (\real_of_int n * u - s\) = real_of_int j - real_of_int \s\ ))" + by (simp only: algebra_simps of_int_diff[symmetric] of_int_eq_iff) + also have "\ = ((\ j\ {0 .. (n - 1)}. real_of_int n * u - s = real_of_int j - real_of_int \s\ \ real_of_int i rdvd real_of_int n * u - s))" using int_rdvd_iff[where i="i" and t="\real_of_int n * u - s\"] by (auto cong: conj_cong) also have "\ = ?rhs" by(simp cong: conj_cong) (simp add: algebra_simps ) finally show ?thesis . @@ -3820,7 +3797,7 @@ NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\ j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) [0..n - 1]) T)" lemma DVDJ_DVD: - assumes xp:"x\ 0" and x1: "x < 1" and np:"real n > 0" + assumes xp:"x\ 0" and x1: "x < 1" and np:"real_of_int n > 0" shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))" proof- let ?f = "\ j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" @@ -3828,15 +3805,15 @@ from foldr_disj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"] have "Ifm (x#bs) (DVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" by (simp add: np DVDJ_def) - also have "\ = (\ j\ {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \ real i rdvd real (j - floor (- ?s)))" + also have "\ = (\ j\ {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int (floor (- ?s)) + real_of_int j \ real_of_int i rdvd real_of_int (j - floor (- ?s)))" by (simp add: algebra_simps) also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] - have "\ = (real i rdvd real n * x - (-?s))" by simp + have "\ = (real_of_int i rdvd real_of_int n * x - (-?s))" by simp finally show ?thesis by simp qed lemma NDVDJ_NDVD: - assumes xp:"x\ 0" and x1: "x < 1" and np:"real n > 0" + assumes xp:"x\ 0" and x1: "x < 1" and np:"real_of_int n > 0" shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))" proof- let ?f = "\ j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))" @@ -3844,10 +3821,10 @@ from foldr_conj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"] have "Ifm (x#bs) (NDVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" by (simp add: np NDVDJ_def) - also have "\ = (\ (\ j\ {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \ real i rdvd real (j - floor (- ?s))))" + also have "\ = (\ (\ j\ {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int (floor (- ?s)) + real_of_int j \ real_of_int i rdvd real_of_int (j - floor (- ?s))))" by (simp add: algebra_simps) also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] - have "\ = (\ (real i rdvd real n * x - (-?s)))" by simp + have "\ = (\ (real_of_int i rdvd real_of_int n * x - (-?s)))" by simp finally show ?thesis by simp qed @@ -3902,7 +3879,7 @@ moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H DVD_def) } moreover {assume inz: "i\0" and "n<0" hence ?th by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 - rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } + rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)} ultimately show ?th by blast qed @@ -3920,7 +3897,7 @@ moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H NDVD_def) } moreover {assume inz: "i\0" and "n<0" hence ?th by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 - rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } + rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)} ultimately show ?th by blast @@ -4152,16 +4129,16 @@ next case (3 c e) from 3 have nb: "numbound0 e" by simp - from 3 have cp: "real c > 0" by simp + from 3 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real c * x < - ?e)" + hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - hence "real c * x + ?e < 0" by arith - hence "real c * x + ?e \ 0" by simp + hence "real_of_int c * x + ?e < 0" by arith + hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp @@ -4169,16 +4146,16 @@ next case (4 c e) from 4 have nb: "numbound0 e" by simp - from 4 have cp: "real c > 0" by simp + from 4 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real c * x < - ?e)" + hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - hence "real c * x + ?e < 0" by arith - hence "real c * x + ?e \ 0" by simp + hence "real_of_int c * x + ?e < 0" by arith + hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp @@ -4186,15 +4163,15 @@ next case (5 c e) from 5 have nb: "numbound0 e" by simp - from 5 have cp: "real c > 0" by simp + from 5 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real c * x < - ?e)" + hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - hence "real c * x + ?e < 0" by arith + hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp @@ -4202,15 +4179,15 @@ next case (6 c e) from 6 have nb: "numbound0 e" by simp - from 6 have cp: "real c > 0" by simp + from 6 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real c * x < - ?e)" + hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - hence "real c * x + ?e < 0" by arith + hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp @@ -4218,15 +4195,15 @@ next case (7 c e) from 7 have nb: "numbound0 e" by simp - from 7 have cp: "real c > 0" by simp + from 7 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real c * x < - ?e)" + hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - hence "real c * x + ?e < 0" by arith + hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp @@ -4234,15 +4211,15 @@ next case (8 c e) from 8 have nb: "numbound0 e" by simp - from 8 have cp: "real c > 0" by simp + from 8 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real c * x < - ?e)" + hence "(real_of_int c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) - hence "real c * x + ?e < 0" by arith + hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp @@ -4260,16 +4237,16 @@ next case (3 c e) from 3 have nb: "numbound0 e" by simp - from 3 have cp: "real c > 0" by simp + from 3 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - hence "real c * x + ?e > 0" by arith - hence "real c * x + ?e \ 0" by simp + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + hence "real_of_int c * x + ?e > 0" by arith + hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp @@ -4277,16 +4254,16 @@ next case (4 c e) from 4 have nb: "numbound0 e" by simp - from 4 have cp: "real c > 0" by simp + from 4 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - hence "real c * x + ?e > 0" by arith - hence "real c * x + ?e \ 0" by simp + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + hence "real_of_int c * x + ?e > 0" by arith + hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp @@ -4294,15 +4271,15 @@ next case (5 c e) from 5 have nb: "numbound0 e" by simp - from 5 have cp: "real c > 0" by simp + from 5 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - hence "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp @@ -4310,15 +4287,15 @@ next case (6 c e) from 6 have nb: "numbound0 e" by simp - from 6 have cp: "real c > 0" by simp + from 6 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - hence "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp @@ -4326,15 +4303,15 @@ next case (7 c e) from 7 have nb: "numbound0 e" by simp - from 7 have cp: "real c > 0" by simp + from 7 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - hence "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp @@ -4342,15 +4319,15 @@ next case (8 c e) from 8 have nb: "numbound0 e" by simp - from 8 have cp: "real c > 0" by simp + from 8 have cp: "real_of_int c > 0" by simp fix a let ?e="Inum (a#bs) e" - let ?z = "(- ?e) / real c" + let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp - have "(real c * x > - ?e)" by (simp add: ac_simps) - hence "real c * x + ?e > 0" by arith + have "(real_of_int c * x > - ?e)" by (simp add: ac_simps) + hence "real_of_int c * x + ?e > 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp @@ -4423,78 +4400,78 @@ "\ p = (\ (t,n). p)" lemma \_I: assumes lp: "isrlfm p" - and np: "real n > 0" and nbt: "numbound0 t" - shows "(Ifm (x#bs) (\ p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \ bound0 (\ p (t,n))" (is "(?I x (\ p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") + and np: "real_of_int n > 0" and nbt: "numbound0 t" + shows "(Ifm (x#bs) (\ p (t,n)) = Ifm (((Inum (x#bs) t)/(real_of_int n))#bs) p) \ bound0 (\ p (t,n))" (is "(?I x (\ p (t,n)) = ?I ?u p) \ ?B p" is "(_ = ?I (?t/?n) p) \ _" is "(_ = ?I (?N x t /_) p) \ _") using lp proof(induct p rule: \.induct) case (5 c e) from 5 have cp: "c >0" and nb: "numbound0 e" by simp_all - have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)" + have "?I ?u (Lt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) < 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)" - by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) < 0)" + by (simp only: pos_less_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) < 0)" + also have "\ = (real_of_int c *?t + ?n* (?N x e) < 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (6 c e) from 6 have cp: "c >0" and nb: "numbound0 e" by simp_all - have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + have "?I ?u (Le (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_le_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (7 c e) from 7 have cp: "c >0" and nb: "numbound0 e" by simp_all - have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)" + have "?I ?u (Gt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) > 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)" - by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) > 0)" + by (simp only: pos_divide_less_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) > 0)" + also have "\ = (real_of_int c *?t + ?n* (?N x e) > 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (8 c e) from 8 have cp: "c >0" and nb: "numbound0 e" by simp_all - have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + have "?I ?u (Ge (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: pos_divide_le_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (3 c e) from 3 have cp: "c >0" and nb: "numbound0 e" by simp_all - from np have np: "real n \ 0" by simp - have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)" + from np have np: "real_of_int n \ 0" by simp + have "?I ?u (Eq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) = 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) = 0)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) = 0)" + also have "\ = (real_of_int c *?t + ?n* (?N x e) = 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (4 c e) from 4 have cp: "c >0" and nb: "numbound0 e" by simp_all - from np have np: "real n \ 0" by simp - have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \ 0)" + from np have np: "real_of_int n \ 0" by simp + have "?I ?u (NEq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp - also have "\ = (?n*(real c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" + also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) - also have "\ = (real c *?t + ?n* (?N x e) \ 0)" + also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) -qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"]) +qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real_of_int n" and b'="x"]) lemma \_l: assumes lp: "isrlfm p" @@ -4506,14 +4483,14 @@ assumes lp: "isrlfm p" and nmi: "\ (Ifm (a#bs) (minusinf p))" (is "\ (Ifm (a#bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") + shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real_of_int m" (is "\ (s,m) \ ?U p. x \ ?N a s / real_of_int m") proof- - have "\ (s,m) \ set (\ p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") + have "\ (s,m) \ set (\ p). real_of_int m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real_of_int m *x \ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) - then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast - from \_l[OF lp] smU have mp: "real m > 0" by auto - from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real_of_int m * x \ ?N a s" by blast + from \_l[OF lp] smU have mp: "real_of_int m > 0" by auto + from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) thus ?thesis using smU by auto qed @@ -4522,119 +4499,119 @@ assumes lp: "isrlfm p" and nmi: "\ (Ifm (a#bs) (plusinf p))" (is "\ (Ifm (a#bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") - shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real m" (is "\ (s,m) \ ?U p. x \ ?N a s / real m") + shows "\ (s,m) \ set (\ p). x \ Inum (a#bs) s / real_of_int m" (is "\ (s,m) \ ?U p. x \ ?N a s / real_of_int m") proof- - have "\ (s,m) \ set (\ p). real m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real m *x \ ?N a s") + have "\ (s,m) \ set (\ p). real_of_int m * x \ Inum (a#bs) s " (is "\ (s,m) \ ?U p. real_of_int m *x \ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) - then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real m * x \ ?N a s" by blast - from \_l[OF lp] smU have mp: "real m > 0" by auto - from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real m" + then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real_of_int m * x \ ?N a s" by blast + from \_l[OF lp] smU have mp: "real_of_int m > 0" by auto + from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) thus ?thesis using smU by auto qed lemma lin_dense: assumes lp: "isrlfm p" - and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real n) ` set (\ p)" - (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real n ) ` (?U p)") + and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real_of_int n) ` set (\ p)" + (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real_of_int n ) ` (?U p)") and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" and ly: "l < y" and yu: "y < u" shows "Ifm (y#bs) p" using lp px noS proof (induct p rule: isrlfm.induct) - case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all - from 5 have "x * real c + ?N x e < 0" by (simp add: algebra_simps) - hence pxc: "x < (- ?N x e) / real c" + case (5 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all + from 5 have "x * real_of_int c + ?N x e < 0" by (simp add: algebra_simps) + hence pxc: "x < (- ?N x e) / real_of_int c" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from 5 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y < (-?N x e)/ real c" - hence "y * real c < - ?N x e" + from 5 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto + hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto + moreover {assume y: "y < (-?N x e)/ real_of_int c" + hence "y * real_of_int c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e < 0" by (simp add: algebra_simps) + hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real c" - with yu have eu: "u > (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) + moreover {assume y: "y > (- ?N x e) / real_of_int c" + with yu have eu: "u > (- ?N x e) / real_of_int c" by auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ l" by (cases "(- ?N x e) / real_of_int c > l", auto) with lx pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next - case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all - from 6 have "x * real c + ?N x e \ 0" by (simp add: algebra_simps) - hence pxc: "x \ (- ?N x e) / real c" + case (6 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all + from 6 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) + hence pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) - from 6 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y < (-?N x e)/ real c" - hence "y * real c < - ?N x e" + from 6 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto + hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto + moreover {assume y: "y < (-?N x e)/ real_of_int c" + hence "y * real_of_int c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e < 0" by (simp add: algebra_simps) + hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real c" - with yu have eu: "u > (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ l" by (cases "(- ?N x e) / real c > l", auto) + moreover {assume y: "y > (- ?N x e) / real_of_int c" + with yu have eu: "u > (- ?N x e) / real_of_int c" by auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ l" by (cases "(- ?N x e) / real_of_int c > l", auto) with lx pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next - case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all - from 7 have "x * real c + ?N x e > 0" by (simp add: algebra_simps) - hence pxc: "x > (- ?N x e) / real c" + case (7 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all + from 7 have "x * real_of_int c + ?N x e > 0" by (simp add: algebra_simps) + hence pxc: "x > (- ?N x e) / real_of_int c" by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) - from 7 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y > (-?N x e)/ real c" - hence "y * real c > - ?N x e" + from 7 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto + hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto + moreover {assume y: "y > (-?N x e)/ real_of_int c" + hence "y * real_of_int c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e > 0" by (simp add: algebra_simps) + hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real c" - with ly have eu: "l < (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) + moreover {assume y: "y < (- ?N x e) / real_of_int c" + with ly have eu: "l < (- ?N x e) / real_of_int c" by auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ u" by (cases "(- ?N x e) / real_of_int c > l", auto) with xu pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next - case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all - from 8 have "x * real c + ?N x e \ 0" by (simp add: algebra_simps) - hence pxc: "x \ (- ?N x e) / real c" + case (8 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all + from 8 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) + hence pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) - from 8 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y < (- ?N x e) / real c \ y > (-?N x e) / real c" by auto - moreover {assume y: "y > (-?N x e)/ real c" - hence "y * real c > - ?N x e" + from 8 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto + hence "y < (- ?N x e) / real_of_int c \ y > (-?N x e) / real_of_int c" by auto + moreover {assume y: "y > (-?N x e)/ real_of_int c" + hence "y * real_of_int c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) - hence "real c * y + ?N x e > 0" by (simp add: algebra_simps) + hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real c" - with ly have eu: "l < (- ?N x e) / real c" by auto - with noSc ly yu have "(- ?N x e) / real c \ u" by (cases "(- ?N x e) / real c > l", auto) + moreover {assume y: "y < (- ?N x e) / real_of_int c" + with ly have eu: "l < (- ?N x e) / real_of_int c" by auto + with noSc ly yu have "(- ?N x e) / real_of_int c \ u" by (cases "(- ?N x e) / real_of_int c > l", auto) with xu pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next - case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all - from cp have cnz: "real c \ 0" by simp - from 3 have "x * real c + ?N x e = 0" by (simp add: algebra_simps) - hence pxc: "x = (- ?N x e) / real c" + case (3 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all + from cp have cnz: "real_of_int c \ 0" by simp + from 3 have "x * real_of_int c + ?N x e = 0" by (simp add: algebra_simps) + hence pxc: "x = (- ?N x e) / real_of_int c" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) - from 3 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with lx xu have yne: "x \ - ?N x e / real c" by auto + from 3 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto + with lx xu have yne: "x \ - ?N x e / real_of_int c" by auto with pxc show ?case by simp next - case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all - from cp have cnz: "real c \ 0" by simp - from 4 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real c" by auto - with ly yu have yne: "y \ - ?N x e / real c" by auto - hence "y* real c \ -?N x e" + case (4 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all + from cp have cnz: "real_of_int c \ 0" by simp + from 4 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto + with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto + hence "y* real_of_int c \ -?N x e" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp - hence "y* real c + ?N x e \ 0" by (simp add: algebra_simps) + hence "y* real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by (simp add: algebra_simps) qed (auto simp add: numbound0_I[where bs="bs" and b="y" and b'="x"]) @@ -4645,7 +4622,7 @@ and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") shows "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). - ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" + ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int m) / 2) p" proof- let ?N = "\ x t. Inum (x#bs) t" let ?U = "set (\ p)" @@ -4654,46 +4631,46 @@ have nmi': "\ (?I a (?M p))" by simp from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi have npi': "\ (?I a (?P p))" by simp - have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" + have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). ?I ((?N a l/real_of_int n + ?N a s /real_of_int m) / 2) p" proof- - let ?M = "(\ (t,c). ?N a t / real c) ` ?U" + let ?M = "(\ (t,c). ?N a t / real_of_int c) ` ?U" have fM: "finite ?M" by auto from rminusinf_\[OF lp nmi pa] rplusinf_\[OF lp npi pa] - have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). a \ ?N x l / real n \ a \ ?N x s / real m" by blast + have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). a \ ?N x l / real_of_int n \ a \ ?N x s / real_of_int m" by blast then obtain "t" "n" "s" "m" where tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" - and xs1: "a \ ?N x s / real m" and tx1: "a \ ?N x t / real n" by blast - from \_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real m" and tx: "a \ ?N a t / real n" by auto + and xs1: "a \ ?N x s / real_of_int m" and tx1: "a \ ?N x t / real_of_int n" by blast + from \_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real_of_int m" and tx: "a \ ?N a t / real_of_int n" by auto from tnU have Mne: "?M \ {}" by auto hence Une: "?U \ {}" by simp let ?l = "Min ?M" let ?u = "Max ?M" have linM: "?l \ ?M" using fM Mne by simp have uinM: "?u \ ?M" using fM Mne by simp - have tnM: "?N a t / real n \ ?M" using tnU by auto - have smM: "?N a s / real m \ ?M" using smU by auto + have tnM: "?N a t / real_of_int n \ ?M" using tnU by auto + have smM: "?N a s / real_of_int m \ ?M" using smU by auto have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto - have "?l \ ?N a t / real n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp - have "?N a s / real m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp + have "?l \ ?N a t / real_of_int n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp + have "?N a s / real_of_int m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] have "(\ s\ ?M. ?I s p) \ (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . moreover { fix u assume um: "u\ ?M" and pu: "?I u p" - hence "\ (tu,nu) \ ?U. u = ?N a tu / real nu" by auto - then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real nu" by blast + hence "\ (tu,nu) \ ?U. u = ?N a tu / real_of_int nu" by auto + then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real_of_int nu" by blast have "(u + u) / 2 = u" by auto with pu tuu - have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp + have "?I (((?N a tu / real_of_int nu) + (?N a tu / real_of_int nu)) / 2) p" by simp with tuU have ?thesis by blast} moreover{ assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" by blast - from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real t1n" by auto - then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast - from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real t2n" by auto - then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast + from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real_of_int t1n" by auto + then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \ ?U" and t1u: "t1 = ?N a t1u / real_of_int t1n" by blast + from t2M have "\ (t2u,t2n) \ ?U. t2 = ?N a t2u / real_of_int t2n" by auto + then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \ ?U" and t2u: "t2 = ?N a t2u / real_of_int t2n" by blast from t1x xt2 have t1t2: "t1 < t2" by simp let ?u = "(t1 + t2) / 2" from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto @@ -4702,11 +4679,11 @@ ultimately show ?thesis by blast qed then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" - and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast + and pu: "?I ((?N a l / real_of_int n + ?N a s / real_of_int m) / 2) p" by blast from lnU smU \_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu - have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp + have "?I ((?N x l / real_of_int n + ?N x s / real_of_int m) / 2) p" by simp with lnU smU show ?thesis by auto qed @@ -4714,7 +4691,7 @@ theorem fr_eq: assumes lp: "isrlfm p" - shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (\ p). \ (s,m) \ set (\ p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))" + shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (\ p). \ (s,m) \ set (\ p). Ifm ((((Inum (x#bs) t)/ real_of_int n + (Inum (x#bs) s) / real_of_int m) /2)#bs) p))" (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") proof assume px: "\ x. ?I x p" @@ -4741,18 +4718,18 @@ have "?M \ ?P \ (\ ?M \ \ ?P)" by blast moreover {assume "?M \ ?P" hence "?D" by blast} moreover {assume nmi: "\ ?M" and npi: "\ ?P" - let ?f ="\ (t,n). Inum (x#bs) t / real n" + let ?f ="\ (t,n). Inum (x#bs) t / real_of_int n" let ?N = "\ t. Inum (x#bs) t" {fix t n s m assume "(t,n)\ set (\ p)" and "(s,m) \ set (\ p)" - with \_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" + with \_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute) + from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnp mp np by (simp add: algebra_simps add_divide_distrib) from \_I[OF lp mnp st_nb, where x="x" and bs="bs"] - have "?I x (\ p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])} + have "?I x (\ p (?st,2*n*m)) = ?I ((?N t / real_of_int n + ?N s / real_of_int m) /2) p" by (simp only: st[symmetric])} with rinf_\[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} ultimately show "?D" by blast next @@ -4761,9 +4738,9 @@ moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } moreover {fix t k s l assume "(t,k) \ set (\ p)" and "(s,l) \ set (\ p)" and px:"?I x (\ p (Add (Mul l t) (Mul k s), 2*k*l))" - with \_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto + with \_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int k > 0" and snb: "numbound0 s" and mp:"real_of_int l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" - from np mp have mnp: "real (2*k*l) > 0" by (simp add: mult.commute) + from np mp have mnp: "real_of_int (2*k*l) > 0" by (simp add: mult.commute) from tnb snb have st_nb: "numbound0 ?st" by simp from \_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} ultimately show "?E" by blast @@ -4772,17 +4749,17 @@ text\The overall Part\ lemma real_ex_int_real01: - shows "(\ (x::real). P x) = (\ (i::int) (u::real). 0\ u \ u< 1 \ P (real i + u))" + shows "(\ (x::real). P x) = (\ (i::int) (u::real). 0\ u \ u< 1 \ P (real_of_int i + u))" proof(auto) fix x assume Px: "P x" let ?i = "floor x" - let ?u = "x - real ?i" - have "x = real ?i + ?u" by simp - hence "P (real ?i + ?u)" using Px by simp - moreover have "real ?i \ x" using real_of_int_floor_le by simp hence "0 \ ?u" by arith + let ?u = "x - real_of_int ?i" + have "x = real_of_int ?i + ?u" by simp + hence "P (real_of_int ?i + ?u)" using Px by simp + moreover have "real_of_int ?i \ x" using of_int_floor_le by simp hence "0 \ ?u" by arith moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith - ultimately show "(\ (i::int) (u::real). 0\ u \ u< 1 \ P (real i + u))" by blast + ultimately show "(\ (i::int) (u::real). 0\ u \ u< 1 \ P (real_of_int i + u))" by blast qed fun exsplitnum :: "num \ num" where @@ -4826,11 +4803,11 @@ lemma splitex: assumes qf: "qfree p" - shows "(Ifm bs (E p)) = (\ (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs") + shows "(Ifm bs (E p)) = (\ (i::int). Ifm (real_of_int i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs") proof- - have "?rhs = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm (x#(real i)#bs) (exsplit p))" + have "?rhs = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm (x#(real_of_int i)#bs) (exsplit p))" by (simp add: myless[of _ "1"] myless[of _ "0"] ac_simps) - also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real i + x) #bs) p)" + also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real_of_int i + x) #bs) p)" by (simp only: exsplit[OF qf] ac_simps) also have "\ = (\ x. Ifm (x#bs) p)" by (simp only: real_ex_int_real01[where P="\ x. Ifm (x#bs) p"]) @@ -4880,10 +4857,10 @@ then obtain t n s m where aU:"(t,n) \ ?U" and bU:"(s,m)\ ?U" and rqx: "?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] by (auto simp add: rsplit_def lt_def ge_def) - from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by (auto simp add: split_def) + from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0" by (auto simp add: split_def) let ?st = "Add (Mul m t) (Mul n s)" from tnb snb have stnb: "numbound0 ?st" by simp - from np mp have mnp: "real (2*n*m) > 0" by (simp add: mult.commute) + from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute) from conjunct1[OF \_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx have "\ x. ?I x ?rq" by auto thus "?E" @@ -4894,7 +4871,7 @@ lemma \_cong_aux: assumes Ul: "\ (t,n) \ set U. numbound0 t \ n >0" - shows "((\ (t,n). Inum (x#bs) t /real n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \ set U))" + shows "((\ (t,n). Inum (x#bs) t /real_of_int n) ` (set (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real_of_int n + Inum (x#bs) s /real_of_int m)/2) ` (set U \ set U))" (is "?lhs = ?rhs") proof(auto) fix t n s m @@ -4905,13 +4882,13 @@ let ?st= "Add (Mul m t) (Mul n s)" from Ul th have mnz: "m \ 0" by auto from Ul th have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) - thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / - (2 * real n * real m) + thus "(real_of_int m * Inum (x # bs) t + real_of_int n * Inum (x # bs) s) / + (2 * real_of_int n * real_of_int m) \ (\((t, n), s, m). - (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` + (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) ` (set U \ set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) by (rule_tac x="(s,m)" in bexI,simp_all) @@ -4923,9 +4900,9 @@ let ?st= "Add (Mul m t) (Mul n s)" from Ul smU have mnz: "m \ 0" by auto from Ul tnU have nnz: "n \ 0" by auto - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) - let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" + let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" have Pc:"\ a b. ?P a b = ?P b a" by auto from Ul alluopairs_set1 have Up:"\ ((t,n),(s,m)) \ set (alluopairs U). n \ 0 \ m \ 0" by blast @@ -4936,13 +4913,13 @@ and Pts': "?P (t',n') (s',m')" by blast from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto let ?st' = "Add (Mul m' t') (Mul n' s')" - have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" + have st': "(?N t' / real_of_int n' + ?N s' / real_of_int m')/2 = ?N ?st' / real_of_int (2*n'*m')" using mnz' nnz' by (simp add: algebra_simps add_divide_distrib) from Pts' have - "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp - also have "\ = ((\(t, n). Inum (x # bs) t / real n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') - finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 - \ (\(t, n). Inum (x # bs) t / real n) ` + "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" by simp + also have "\ = ((\(t, n). Inum (x # bs) t / real_of_int n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') + finally show "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2 + \ (\(t, n). Inum (x # bs) t / real_of_int n) ` (\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` set (alluopairs U)" using ts'_U by blast @@ -4950,7 +4927,7 @@ lemma \_cong: assumes lp: "isrlfm p" - and UU': "((\ (t,n). Inum (x#bs) t /real n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") + and UU': "((\ (t,n). Inum (x#bs) t /real_of_int n) ` U') = ((\ ((t,n),(s,m)). (Inum (x#bs) t /real_of_int n + Inum (x#bs) s /real_of_int m)/2) ` (U \ U))" (is "?f ` U' = ?g ` (U\U)") and U: "\ (t,n) \ U. numbound0 t \ n > 0" and U': "\ (t,n) \ U'. numbound0 t \ n > 0" shows "(\ (t,n) \ U. \ (s,m) \ U. Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))) = (\ (t,n) \ U'. Ifm (x#bs) (\ p (t,n)))" @@ -4963,18 +4940,18 @@ from tnU smU U have tnb: "numbound0 t" and np: "n > 0" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2*n*m) > 0" - by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) + from np mp have mnp: "real_of_int (2*n*m) > 0" + by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) from tnU smU UU' have "?g ((t,n),(s,m)) \ ?f ` U'" by blast hence "\ (t',n') \ U'. ?g ((t,n),(s,m)) = ?f (t',n')" by auto (rule_tac x="(a,b)" in bexI, auto) then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto + from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto from \_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp from conjunct1[OF \_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]] have "Ifm (x # bs) (\ p (t', n')) " by (simp only: st) then show ?rhs using tnU' by auto @@ -4991,14 +4968,14 @@ from tnU smU U have tnb: "numbound0 t" and np: "n > 0" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real (2*n*m) > 0" - by (simp add: mult.commute real_of_int_mult[symmetric] del: real_of_int_mult) + from np mp have mnp: "real_of_int (2*n*m) > 0" + by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp - have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" + have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mp np by (simp add: algebra_simps add_divide_distrib) - from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto + from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto from \_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' - have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp + have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp with \_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast qed @@ -5016,8 +4993,8 @@ let ?S = "map ?g ?Up" let ?SS = "map simp_num_pair ?S" let ?Y = "remdups ?SS" - let ?f= "(\ (t,n). ?N t / real n)" - let ?h = "\ ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" + let ?f= "(\ (t,n). ?N t / real_of_int n)" + let ?h = "\ ((t,n),(s,m)). (?N t/real_of_int n + ?N s/ real_of_int m) /2" let ?F = "\ p. \ a \ set (\ p). \ b \ set (\ p). ?I x (\ p (?g(a,b)))" let ?ep = "evaldjf (\ ?q) ?Y" from rlfm_l[OF qf] have lq: "isrlfm ?q" @@ -5058,7 +5035,7 @@ have "\ (t,n) \ set ?Y. bound0 (\ ?q (t,n))" proof- { fix t n assume tnY: "(t,n) \ set ?Y" - with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto + with Y_l have tnb: "numbound0 t" and np: "real_of_int n > 0" by auto from \_I[OF lq np tnb] have "bound0 (\ ?q (t,n))" by simp} thus ?thesis by blast @@ -5080,9 +5057,9 @@ qed lemma cp_thm': - assumes lp: "iszlfm p (real (i::int)#bs)" + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" and up: "d_\ p 1" and dd: "d_\ p d" and dp: "d > 0" - shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. d}. Ifm (real j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real i#bs)) ` set (\ p). Ifm ((b+real j)#bs) p))" + shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. d}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real_of_int i#bs)) ` set (\ p). Ifm ((b+real_of_int j)#bs) p))" using cp_thm[OF lp up dd dp] by auto definition unit :: "fm \ fm \ num list \ int" where @@ -5091,14 +5068,18 @@ in if length B \ length a then (q,B,d) else (mirror q, a,d))" lemma unit: assumes qf: "qfree p" - shows "\ q B d. unit p = (q,B,d) \ ((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ (Inum (real i#bs)) ` set B = (Inum (real i#bs)) ` set (\ q) \ d_\ q 1 \ d_\ q d \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ b\ set B. numbound0 b)" + shows "\ q B d. unit p = (q,B,d) \ + ((\ (x::int). Ifm (real_of_int x#bs) p) = + (\ (x::int). Ifm (real_of_int x#bs) q)) \ + (Inum (real_of_int i#bs)) ` set B = (Inum (real_of_int i#bs)) ` set (\ q) \ + d_\ q 1 \ d_\ q d \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ b\ set B. numbound0 b)" proof- fix q B d assume qBd: "unit p = (q,B,d)" - let ?thes = "((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ - Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\ q) \ - d_\ q 1 \ d_\ q d \ 0 < d \ iszlfm q (real i # bs) \ (\ b\ set B. numbound0 b)" - let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = (\ (x::int). Ifm (real_of_int x#bs) q)) \ + Inum (real_of_int i#bs) ` set B = Inum (real_of_int i#bs) ` set (\ q) \ + d_\ q 1 \ d_\ q d \ 0 < d \ iszlfm q (real_of_int i # bs) \ (\ b\ set B. numbound0 b)" + let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?p' = "zlfm p" let ?l = "\ ?p'" let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a_\ ?p' ?l)" @@ -5110,20 +5091,20 @@ from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] have pp': "\ i. ?I i ?p' = ?I i p" by auto from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]] - have lp': "\ (i::int). iszlfm ?p' (real i#bs)" by simp - hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp + have lp': "\ (i::int). iszlfm ?p' (real_of_int i#bs)" by simp + hence lp'': "iszlfm ?p' (real_of_int (i::int)#bs)" by simp from lp' \[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d_\ ?p' ?l" by auto from a_\_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp' have pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by (simp add: int_rdvd_iff) - from lp'' lp a_\[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d_\ ?q 1" + from lp'' lp a_\[OF lp'' dl lp] have lq:"iszlfm ?q (real_of_int i#bs)" and uq: "d_\ ?q 1" by (auto simp add: isint_def) from \[OF lq] have dp:"?d >0" and dd: "d_\ ?q ?d" by blast+ - let ?N = "\ t. Inum (real (i::int)#bs) t" + let ?N = "\ t. Inum (real_of_int (i::int)#bs) t" have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_comp) - also have "\ = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto + also have "\ = ?N ` ?B" using simpnum_ci[where bs="real_of_int i #bs"] by auto finally have BB': "?N ` set ?B' = ?N ` ?B" . have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_comp) - also have "\ = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto + also have "\ = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] by auto finally have AA': "?N ` set ?A' = ?N ` ?A" . from \_numbound0[OF lq] have B_nb:"\ b\ set ?B'. numbound0 b" by simp @@ -5143,8 +5124,8 @@ and bn: "\b\ set B. numbound0 b" by simp+ from mirror_ex[OF lq] pq_ex q have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp - from lq uq q mirror_d_\ [where p="?q" and bs="bs" and a="real i"] - have lq': "iszlfm q (real i#bs)" and uq: "d_\ q 1" by auto + from lq uq q mirror_d_\ [where p="?q" and bs="bs" and a="real_of_int i"] + have lq': "iszlfm q (real_of_int i#bs)" and uq: "d_\ q 1" by auto from \[OF lq'] mirror_\[OF lq] q d have dq:"d_\ q d " by auto from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp } @@ -5163,11 +5144,11 @@ [(b,j). b\B,j\js])) in decr (disj md qd)))" lemma cooper: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" + shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" (is "(?lhs = ?rhs) \ _") proof- - let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "fst (unit p)" let ?B = "fst (snd(unit p))" let ?d = "snd (snd (unit p))" @@ -5176,7 +5157,7 @@ let ?smq = "simpfm ?mq" let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" fix i - let ?N = "\ t. Inum (real (i::int)#bs) t" + let ?N = "\ t. Inum (real_of_int (i::int)#bs) t" let ?bjs = "[(b,j). b\?B,j\?js]" let ?sbjs = "map (\ (b,j). simpnum (Add b (C j))) ?bjs" let ?qd = "evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs)" @@ -5184,7 +5165,7 @@ from unit[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and B:"?N ` set ?B = ?N ` set (\ ?q)" and uq:"d_\ ?q 1" and dd: "d_\ ?q ?d" and dp: "?d > 0" and - lq: "iszlfm ?q (real i#bs)" and + lq: "iszlfm ?q (real_of_int i#bs)" and Bn: "\ b\ set ?B. numbound0 b" by auto from zlin_qfree[OF lq] have qfq: "qfree ?q" . from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". @@ -5207,8 +5188,8 @@ from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B - have "?lhs = (\ j\ {1.. ?d}. ?I j ?mq \ (\ b\ ?N ` set ?B. Ifm ((b+ real j)#bs) ?q))" by auto - also have "\ = ((\ j\ set ?js. ?I j ?smq) \ (\ (b,j) \ (?N ` set ?B \ set ?js). Ifm ((b+ real j)#bs) ?q))" by auto + have "?lhs = (\ j\ {1.. ?d}. ?I j ?mq \ (\ b\ ?N ` set ?B. Ifm ((b+ real_of_int j)#bs) ?q))" by auto + also have "\ = ((\ j\ set ?js. ?I j ?smq) \ (\ (b,j) \ (?N ` set ?B \ set ?js). Ifm ((b+ real_of_int j)#bs) ?q))" by auto also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (simpnum (Add b (C j)))) ` set ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci) also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ set ?sbjs. Ifm (?N t #bs) ?q))" @@ -5233,15 +5214,15 @@ lemma DJcooper: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \ qfree (DJ cooper p)" + shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ cooper p))) \ qfree (DJ cooper p)" proof- from cooper have cqf: "\ p. qfree p \ qfree (cooper p)" by blast from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast have "Ifm bs (DJ cooper p) = (\ q\ set (disjuncts p). Ifm bs (cooper q))" by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real x#bs) q)" + also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" using cooper disjuncts_qf[OF qf] by blast - also have "\ = (\ (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) + also have "\ = (\ (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.induct, auto) finally show ?thesis using thqf by blast qed @@ -5292,9 +5273,9 @@ qed lemma rl_thm': - assumes lp: "iszlfm p (real (i::int)#bs)" + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" and R: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" - shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" + shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp definition chooset :: "fm \ fm \ ((num\int) list) \ int" where @@ -5304,12 +5285,18 @@ in if length B \ length a then (q,B,d) else (mirror q, a,d))" lemma chooset: assumes qf: "qfree p" - shows "\ q B d. chooset p = (q,B,d) \ ((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ ((\(t,k). (Inum (real i#bs) t,k)) ` set B = (\(t,k). (Inum (real i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" + shows "\ q B d. chooset p = (q,B,d) \ + ((\ (x::int). Ifm (real_of_int x#bs) p) = + (\ (x::int). Ifm (real_of_int x#bs) q)) \ + ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ + (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" proof- fix q B d assume qBd: "chooset p = (q,B,d)" - let ?thes = "((\ (x::int). Ifm (real x#bs) p) = (\ (x::int). Ifm (real x#bs) q)) \ ((\(t,k). (Inum (real i#bs) t,k)) ` set B = (\(t,k). (Inum (real i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" - let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = + (\ (x::int). Ifm (real_of_int x#bs) q)) \ ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ + (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" + let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "zlfm p" let ?d = "\ ?q" let ?B = "set (\ ?q)" @@ -5320,17 +5307,17 @@ from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] have pp': "\ i. ?I i ?q = ?I i p" by auto hence pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by simp - from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"] - have lq: "iszlfm ?q (real (i::int)#bs)" . + from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real_of_int i"] + have lq: "iszlfm ?q (real_of_int (i::int)#bs)" . from \[OF lq] have dp:"?d >0" by blast - let ?N = "\ (t,c). (Inum (real (i::int)#bs) t,c)" + let ?N = "\ (t,c). (Inum (real_of_int (i::int)#bs) t,c)" have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_comp) also have "\ = ?N ` ?B" - by(simp add: split_def image_comp simpnum_ci[where bs="real i #bs"] image_def) + by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) finally have BB': "?N ` set ?B' = ?N ` ?B" . have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_comp) - also have "\ = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] - by(simp add: split_def image_comp simpnum_ci[where bs="real i #bs"] image_def) + also have "\ = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] + by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) finally have AA': "?N ` set ?A' = ?N ` ?A" . from \_l[OF lq] have B_nb:"\ (e,c)\ set ?B'. numbound0 e \ c > 0" by (simp add: split_def) @@ -5350,8 +5337,8 @@ and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto from mirror_ex[OF lq] pq_ex q have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp - from lq q mirror_l [where p="?q" and bs="bs" and a="real i"] - have lq': "iszlfm q (real i#bs)" by auto + from lq q mirror_l [where p="?q" and bs="bs" and a="real_of_int i"] + have lq': "iszlfm q (real_of_int i#bs)" by auto from mirror_\[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp } ultimately show ?thes by blast @@ -5387,11 +5374,11 @@ in decr (disj md qd)))" lemma redlove: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" + shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" (is "(?lhs = ?rhs) \ _") proof- - let ?I = "\ (x::int) p. Ifm (real x#bs) p" + let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "fst (chooset p)" let ?B = "fst (snd(chooset p))" let ?d = "snd (snd (chooset p))" @@ -5400,12 +5387,12 @@ let ?smq = "simpfm ?mq" let ?md = "evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js" fix i - let ?N = "\ (t,k). (Inum (real (i::int)#bs) t,k)" + let ?N = "\ (t,k). (Inum (real_of_int (i::int)#bs) t,k)" let ?qd = "evaldjf (stage ?q ?d) ?B" have qbf:"chooset p = (?q,?B,?d)" by simp from chooset[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and B:"?N ` set ?B = ?N ` set (\ ?q)" and dd: "\ ?q = ?d" and dp: "?d > 0" and - lq: "iszlfm ?q (real i#bs)" and + lq: "iszlfm ?q (real_of_int i#bs)" and Bn: "\ (e,c)\ set ?B. numbound0 e \ c > 0" by auto from zlin_qfree[OF lq] have qfq: "qfree ?q" . from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". @@ -5420,7 +5407,7 @@ from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) from trans [OF pq_ex rl_thm'[OF lq B]] dd - have "?lhs = ((\ j\ {1.. ?d}. ?I j ?mq) \ (\ (e,c)\ set ?B. \ j\ {1 .. c*?d}. Ifm (real i#bs) (\ ?q c (Add e (C j)))))" by auto + have "?lhs = ((\ j\ {1.. ?d}. ?I j ?mq) \ (\ (e,c)\ set ?B. \ j\ {1 .. c*?d}. Ifm (real_of_int i#bs) (\ ?q c (Add e (C j)))))" by auto also have "\ = ((\ j\ {1.. ?d}. ?I j ?smq) \ (\ (e,c)\ set ?B. ?I i (stage ?q ?d (e,c) )))" by (simp add: stage split_def) also have "\ = ((\ j\ {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \ ?I i ?qd)" @@ -5443,15 +5430,15 @@ lemma DJredlove: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \ qfree (DJ redlove p)" + shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ redlove p))) \ qfree (DJ redlove p)" proof- from redlove have cqf: "\ p. qfree p \ qfree (redlove p)" by blast from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast have "Ifm bs (DJ redlove p) = (\ q\ set (disjuncts p). Ifm bs (redlove q))" by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real x#bs) q)" + also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" using redlove disjuncts_qf[OF qf] by blast - also have "\ = (\ (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) + also have "\ = (\ (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.induct, auto) finally show ?thesis using thqf by blast qed @@ -5473,9 +5460,9 @@ show "qfree (mircfr p)\(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") proof- let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" - have "?rhs = (\ (i::int). \ x. Ifm (x#real i#bs) ?es)" + have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" using splitex[OF qf] by simp - with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ + with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def) qed qed @@ -5487,9 +5474,9 @@ show "qfree (mirlfr p)\(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") proof- let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" - have "?rhs = (\ (i::int). \ x. Ifm (x#real i#bs) ?es)" + have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" using splitex[OF qf] by simp - with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ + with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def) qed qed @@ -5542,8 +5529,8 @@ fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t of NONE => error "Variable not found in the list!" | SOME n => mk_Bound n) - | num_of_term vs @{term "real (0::int)"} = mk_C 0 - | num_of_term vs @{term "real (1::int)"} = mk_C 1 + | num_of_term vs @{term "of_int (0::int)"} = mk_C 0 + | num_of_term vs @{term "of_int (1::int)"} = mk_C 1 | num_of_term vs @{term "0::real"} = mk_C 0 | num_of_term vs @{term "1::real"} = mk_C 1 | num_of_term vs @{term "- 1::real"} = mk_C (~ 1) @@ -5557,13 +5544,13 @@ (case (num_of_term vs t1) of @{code C} i => @{code Mul} (i, num_of_term vs t2) | _ => error "num_of_term: unsupported Multiplication") - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t')) = + | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t')) = mk_C (HOLogic.dest_num t') - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t')) = + | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t')) = mk_C (~ (HOLogic.dest_num t')) - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ t')) = + | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "floor :: real \ int"} $ t')) = @{code Floor} (num_of_term vs t') - | num_of_term vs (@{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ t')) = + | num_of_term vs (@{term "of_int :: int \ real"} $ (@{term "ceiling :: real \ int"} $ t')) = @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) | num_of_term vs (@{term "numeral :: _ \ real"} $ t') = mk_C (HOLogic.dest_num t') @@ -5579,9 +5566,9 @@ @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) - | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t1)) $ t2) = + | fm_of_term vs (@{term "op rdvd"} $ (@{term "of_int :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t1)) $ t2) = mk_Dvd (HOLogic.dest_num t1, num_of_term vs t2) - | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t1)) $ t2) = + | fm_of_term vs (@{term "op rdvd"} $ (@{term "of_int :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t1)) $ t2) = mk_Dvd (~ (HOLogic.dest_num t1), num_of_term vs t2) | fm_of_term vs (@{term "op = :: bool \ bool \ bool"} $ t1 $ t2) = @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) @@ -5599,14 +5586,14 @@ @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); -fun term_of_num vs (@{code C} i) = @{term "real :: int \ real"} $ +fun term_of_num vs (@{code C} i) = @{term "of_int :: int \ real"} $ HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i) | term_of_num vs (@{code Bound} n) = let val m = @{code integer_of_nat} n; in fst (the (find_first (fn (_, q) => m = q) vs)) end | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) = - @{term "real :: int \ real"} $ (@{term "ceiling :: real \ int"} $ term_of_num vs t') + @{term "of_int :: int \ real"} $ (@{term "ceiling :: real \ int"} $ term_of_num vs t') | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \ real"} $ term_of_num vs t' | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \ real \ real"} $ term_of_num vs t1 $ term_of_num vs t2 @@ -5614,7 +5601,7 @@ term_of_num vs t1 $ term_of_num vs t2 | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \ real \ real"} $ term_of_num vs (@{code C} i) $ term_of_num vs t2 - | term_of_num vs (@{code Floor} t) = @{term "real :: int \ real"} $ (@{term "floor :: real \ int"} $ term_of_num vs t) + | term_of_num vs (@{code Floor} t) = @{term "of_int :: int \ real"} $ (@{term "floor :: real \ int"} $ term_of_num vs t) | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)) | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s)); @@ -5665,12 +5652,11 @@ Scan.lift (Args.mode "no_quantify") >> (fn q => fn ctxt => SIMPLE_METHOD' (Mir_Tac.mir_tac ctxt (not q))) \ "decision procedure for MIR arithmetic" - - -lemma "\x::real. (\x\ = \x\ = (x = real \x\))" +(*FIXME +lemma "\x::real. (\x\ = \x\ \ (x = real_of_int \x\))" by mir -lemma "\x::real. real (2::int)*x - (real (1::int)) < real \x\ + real \x\ \ real \x\ + real \x\ \ real (2::int)*x + (real (1::int))" +lemma "\x::real. real_of_int (2::int)*x - (real_of_int (1::int)) < real_of_int \x\ + real_of_int \x\ \ real_of_int \x\ + real_of_int \x\ \ real_of_int (2::int)*x + (real_of_int (1::int))" by mir lemma "\x::real. 2*\x\ \ \2*x\ \ \2*x\ \ 2*\x+1\" @@ -5681,6 +5667,6 @@ lemma "\(x::real) (y::real). \x\ = \y\ \ 0 \ abs (y - x) \ abs (y - x) \ 1" by mir - +*) end diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Decision_Procs/approximation_generator.ML --- a/src/HOL/Decision_Procs/approximation_generator.ML Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Decision_Procs/approximation_generator.ML Tue Nov 10 14:18:41 2015 +0000 @@ -94,17 +94,17 @@ val postproc_form_eqs = @{lemma - "real (Float 0 a) = 0" - "real (Float (numeral m) 0) = numeral m" - "real (Float 1 0) = 1" - "real (Float (- 1) 0) = - 1" - "real (Float 1 (numeral e)) = 2 ^ numeral e" - "real (Float 1 (- numeral e)) = 1 / 2 ^ numeral e" - "real (Float a 1) = a * 2" - "real (Float a (-1)) = a / 2" - "real (Float (- a) b) = - real (Float a b)" - "real (Float (numeral m) (numeral e)) = numeral m * 2 ^ (numeral e)" - "real (Float (numeral m) (- numeral e)) = numeral m / 2 ^ (numeral e)" + "real_of_float (Float 0 a) = 0" + "real_of_float (Float (numeral m) 0) = numeral m" + "real_of_float (Float 1 0) = 1" + "real_of_float (Float (- 1) 0) = - 1" + "real_of_float (Float 1 (numeral e)) = 2 ^ numeral e" + "real_of_float (Float 1 (- numeral e)) = 1 / 2 ^ numeral e" + "real_of_float (Float a 1) = a * 2" + "real_of_float (Float a (-1)) = a / 2" + "real_of_float (Float (- a) b) = - real_of_float (Float a b)" + "real_of_float (Float (numeral m) (numeral e)) = numeral m * 2 ^ (numeral e)" + "real_of_float (Float (numeral m) (- numeral e)) = numeral m / 2 ^ (numeral e)" "- (c * d::real) = -c * d" "- (c / d::real) = -c / d" "- (0::real) = 0" @@ -137,7 +137,7 @@ val ts' = map (AList.lookup op = (map dest_Free xs ~~ ts) #> the_default Term.dummy - #> curry op $ @{term "real::float\_"} + #> curry op $ @{term "real_of_float::float\_"} #> conv_term ctxt (rewrite_with ctxt postproc_form_eqs)) frees in diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Decision_Procs/ferrack_tac.ML --- a/src/HOL/Decision_Procs/ferrack_tac.ML Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Decision_Procs/ferrack_tac.ML Tue Nov 10 14:18:41 2015 +0000 @@ -10,8 +10,8 @@ structure Ferrack_Tac: FERRACK_TAC = struct -val ferrack_ss = let val ths = [@{thm real_of_int_inject}, @{thm real_of_int_less_iff}, - @{thm real_of_int_le_iff}] +val ferrack_ss = let val ths = [@{thm of_int_eq_iff}, @{thm of_int_less_iff}, + @{thm of_int_le_iff}] in @{context} delsimps ths addsimps (map (fn th => th RS sym) ths) end |> simpset_of; diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Decision_Procs/mir_tac.ML --- a/src/HOL/Decision_Procs/mir_tac.ML Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Decision_Procs/mir_tac.ML Tue Nov 10 14:18:41 2015 +0000 @@ -11,7 +11,7 @@ struct val mir_ss = -let val ths = [@{thm "real_of_int_inject"}, @{thm "real_of_int_less_iff"}, @{thm "real_of_int_le_iff"}] +let val ths = [@{thm "of_int_eq_iff"}, @{thm "of_int_less_iff"}, @{thm "of_int_le_iff"}] in simpset_of (@{context} delsimps ths addsimps (map (fn th => th RS sym) ths)) end; @@ -23,9 +23,9 @@ (map (fn th => th RS sym) [@{thm "numeral_1_eq_1"}]) @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps} val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, - @{thm real_of_nat_numeral}, - @{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"}, - @{thm "real_of_int_zero"}, @{thm "real_of_nat_zero"}, + @{thm of_nat_numeral}, + @{thm "of_nat_Suc"}, @{thm "of_nat_1"}, + @{thm "of_int_0"}, @{thm "of_nat_0"}, @{thm "divide_zero"}, @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_right"}, @{thm "times_divide_eq_left"}, @{thm "divide_divide_eq_right"}, diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Deriv.thy Tue Nov 10 14:18:41 2015 +0000 @@ -841,9 +841,7 @@ (auto simp: has_field_derivative_def) lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" - apply (cut_tac DERIV_power [OF DERIV_ident]) - apply (simp add: real_of_nat_def) - done + using DERIV_power [OF DERIV_ident] by simp lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ ((\x. g (f x)) has_field_derivative E * D) (at x within s)" @@ -881,9 +879,6 @@ shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis UNIV_I DERIV_chain_s [of UNIV] assms) -declare - DERIV_power[where 'a=real, unfolded real_of_nat_def[symmetric], derivative_intros] - text\Alternative definition for differentiability\ lemma DERIV_LIM_iff: diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Inequalities.thy --- a/src/HOL/Inequalities.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Inequalities.thy Tue Nov 10 14:18:41 2015 +0000 @@ -31,9 +31,10 @@ have "m*(m-1) \ n*(n + 1)" using assms by (meson diff_le_self order_trans le_add1 mult_le_mono) hence "int(\ {m..n}) = int((n*(n+1) - m*(m-1)) div 2)" using assms - by (auto simp: Setsum_Icc_int[transferred, OF assms] zdiv_int int_mult - split: zdiff_int_split) - thus ?thesis by simp + by (auto simp: Setsum_Icc_int[transferred, OF assms] zdiv_int int_mult simp del: of_nat_setsum + split: zdiff_int_split) + thus ?thesis + by blast qed lemma Setsum_Ico_nat: assumes "(m::nat) \ n" @@ -75,7 +76,6 @@ (\i j. \ i\j; j \ b i \ b j) \ n * (\i=0.. (\i=0..i=0.. z \ of_nat (nat z) = of_int z" +lemma of_nat_nat [simp]: "0 \ z \ of_nat (nat z) = of_int z" by transfer (clarsimp simp add: of_nat_diff) end diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Library/Extended_Real.thy Tue Nov 10 14:18:41 2015 +0000 @@ -249,21 +249,15 @@ shows "-a = -b \ a = b" by (cases rule: ereal2_cases[of a b]) simp_all -instantiation ereal :: real_of -begin - -function real_ereal :: "ereal \ real" where - "real_ereal (ereal r) = r" -| "real_ereal \ = 0" -| "real_ereal (-\) = 0" +function real_of_ereal :: "ereal \ real" where + "real_of_ereal (ereal r) = r" +| "real_of_ereal \ = 0" +| "real_of_ereal (-\) = 0" by (auto intro: ereal_cases) termination by standard (rule wf_empty) -instance .. -end - lemma real_of_ereal[simp]: - "real (- x :: ereal) = - (real x)" + "real_of_ereal (- x :: ereal) = - (real_of_ereal x)" by (cases x) simp_all lemma range_ereal[simp]: "range ereal = UNIV - {\, -\}" @@ -378,7 +372,7 @@ instance ereal :: numeral .. -lemma real_of_ereal_0[simp]: "real (0::ereal) = 0" +lemma real_of_ereal_0[simp]: "real_of_ereal (0::ereal) = 0" unfolding zero_ereal_def by simp lemma abs_ereal_zero[simp]: "\0\ = (0::ereal)" @@ -414,13 +408,13 @@ shows "b + a = c + a \ a = \ \ b = c" using assms by (cases rule: ereal3_cases[of a b c]) auto -lemma ereal_real: "ereal (real x) = (if \x\ = \ then 0 else x)" +lemma ereal_real: "ereal (real_of_ereal x) = (if \x\ = \ then 0 else x)" by (cases x) simp_all lemma real_of_ereal_add: fixes a b :: ereal - shows "real (a + b) = - (if (\a\ = \) \ (\b\ = \) \ (\a\ \ \) \ (\b\ \ \) then real a + real b else 0)" + shows "real_of_ereal (a + b) = + (if (\a\ = \) \ (\b\ = \) \ (\a\ \ \) \ (\b\ \ \) then real_of_ereal a + real_of_ereal b else 0)" by (cases rule: ereal2_cases[of a b]) auto @@ -521,7 +515,7 @@ lemma real_of_ereal_positive_mono: fixes x y :: ereal - shows "0 \ x \ x \ y \ y \ \ \ real x \ real y" + shows "0 \ x \ x \ y \ y \ \ \ real_of_ereal x \ real_of_ereal y" by (cases rule: ereal2_cases[of x y]) auto lemma ereal_MInfty_lessI[intro, simp]: @@ -568,24 +562,24 @@ by (cases rule: ereal2_cases[of a b]) auto lemma ereal_le_real_iff: - "x \ real y \ (\y\ \ \ \ ereal x \ y) \ (\y\ = \ \ x \ 0)" + "x \ real_of_ereal y \ (\y\ \ \ \ ereal x \ y) \ (\y\ = \ \ x \ 0)" by (cases y) auto lemma real_le_ereal_iff: - "real y \ x \ (\y\ \ \ \ y \ ereal x) \ (\y\ = \ \ 0 \ x)" + "real_of_ereal y \ x \ (\y\ \ \ \ y \ ereal x) \ (\y\ = \ \ 0 \ x)" by (cases y) auto lemma ereal_less_real_iff: - "x < real y \ (\y\ \ \ \ ereal x < y) \ (\y\ = \ \ x < 0)" + "x < real_of_ereal y \ (\y\ \ \ \ ereal x < y) \ (\y\ = \ \ x < 0)" by (cases y) auto lemma real_less_ereal_iff: - "real y < x \ (\y\ \ \ \ y < ereal x) \ (\y\ = \ \ 0 < x)" + "real_of_ereal y < x \ (\y\ \ \ \ y < ereal x) \ (\y\ = \ \ 0 < x)" by (cases y) auto lemma real_of_ereal_pos: fixes x :: ereal - shows "0 \ x \ 0 \ real x" by (cases x) auto + shows "0 \ x \ 0 \ real_of_ereal x" by (cases x) auto lemmas real_of_ereal_ord_simps = ereal_le_real_iff real_le_ereal_iff ereal_less_real_iff real_less_ereal_iff @@ -599,15 +593,15 @@ lemma abs_ereal_pos[simp]: "0 \ \x :: ereal\" by (cases x) auto -lemma real_of_ereal_le_0[simp]: "real (x :: ereal) \ 0 \ x \ 0 \ x = \" +lemma real_of_ereal_le_0[simp]: "real_of_ereal (x :: ereal) \ 0 \ x \ 0 \ x = \" by (cases x) auto -lemma abs_real_of_ereal[simp]: "\real (x :: ereal)\ = real \x\" +lemma abs_real_of_ereal[simp]: "\real_of_ereal (x :: ereal)\ = real_of_ereal \x\" by (cases x) auto lemma zero_less_real_of_ereal: fixes x :: ereal - shows "0 < real x \ 0 < x \ x \ \" + shows "0 < real_of_ereal x \ 0 < x \ x \ \" by (cases x) auto lemma ereal_0_le_uminus_iff[simp]: @@ -808,7 +802,7 @@ lemma setsum_real_of_ereal: fixes f :: "'i \ ereal" assumes "\x. x \ S \ \f x\ \ \" - shows "(\x\S. real (f x)) = real (setsum f S)" + shows "(\x\S. real_of_ereal (f x)) = real_of_ereal (setsum f S)" proof - have "\x\S. \r. f x = ereal r" proof @@ -886,12 +880,12 @@ lemma one_not_le_zero_ereal[simp]: "\ (1 \ (0::ereal))" by (simp add: one_ereal_def zero_ereal_def) -lemma real_ereal_1[simp]: "real (1::ereal) = 1" +lemma real_ereal_1[simp]: "real_of_ereal (1::ereal) = 1" unfolding one_ereal_def by simp lemma real_of_ereal_le_1: fixes a :: ereal - shows "a \ 1 \ real a \ 1" + shows "a \ 1 \ real_of_ereal a \ 1" by (cases a) (auto simp: one_ereal_def) lemma abs_ereal_one[simp]: "\1\ = (1::ereal)" @@ -1361,10 +1355,10 @@ lemma real_of_ereal_minus: fixes a b :: ereal - shows "real (a - b) = (if \a\ = \ \ \b\ = \ then 0 else real a - real b)" + shows "real_of_ereal (a - b) = (if \a\ = \ \ \b\ = \ then 0 else real_of_ereal a - real_of_ereal b)" by (cases rule: ereal2_cases[of a b]) auto -lemma real_of_ereal_minus': "\x\ = \ \ \y\ = \ \ real x - real y = real (x - y :: ereal)" +lemma real_of_ereal_minus': "\x\ = \ \ \y\ = \ \ real_of_ereal x - real_of_ereal y = real_of_ereal (x - y :: ereal)" by(subst real_of_ereal_minus) auto lemma ereal_diff_positive: @@ -1411,7 +1405,7 @@ lemma real_of_ereal_inverse[simp]: fixes a :: ereal - shows "real (inverse a) = 1 / real a" + shows "real_of_ereal (inverse a) = 1 / real_of_ereal a" by (cases a) (auto simp: inverse_eq_divide) lemma ereal_inverse[simp]: @@ -2307,7 +2301,7 @@ assumes "\x\ \ \" "(f ---> x) F" shows "eventually (\x. \f x\ \ \) F" proof - - have "(f ---> ereal (real x)) F" + have "(f ---> ereal (real_of_ereal x)) F" using assms by (cases x) auto then have "eventually (\x. f x \ ereal ` UNIV) F" by (rule topological_tendstoD) (auto intro: open_ereal) @@ -2371,7 +2365,7 @@ from \open S\ have "open (ereal -` S)" by (rule ereal_openE) - then obtain e where "e > 0" and e: "\y. dist y (real x) < e \ ereal y \ S" + then obtain e where "e > 0" and e: "\y. dist y (real_of_ereal x) < e \ ereal y \ S" using assms unfolding open_dist by force show thesis proof (intro that subsetI) @@ -2379,7 +2373,7 @@ using \0 < e\ by auto fix y assume "y \ {x - ereal e<.. S" using e[of t] by auto @@ -2404,16 +2398,16 @@ lemma lim_real_of_ereal[simp]: assumes lim: "(f ---> ereal x) net" - shows "((\x. real (f x)) ---> x) net" + shows "((\x. real_of_ereal (f x)) ---> x) net" proof (intro topological_tendstoI) fix S assume "open S" and "x \ S" then have S: "open S" "ereal x \ ereal ` S" by (simp_all add: inj_image_mem_iff) - have "\x. f x \ ereal ` S \ real (f x) \ S" + have "\x. f x \ ereal ` S \ real_of_ereal (f x) \ S" by auto from this lim[THEN topological_tendstoD, OF open_ereal, OF S] - show "eventually (\x. real (f x) \ S) net" + show "eventually (\x. real_of_ereal (f x) \ S) net" by (rule eventually_mono) qed @@ -2425,7 +2419,7 @@ { fix l :: ereal assume "\r. eventually (\x. ereal r < f x) F" - from this[THEN spec, of "real l"] have "l \ \ \ eventually (\x. l < f x) F" + from this[THEN spec, of "real_of_ereal l"] have "l \ \ \ eventually (\x. l < f x) F" by (cases l) (auto elim: eventually_elim1) } then show ?thesis @@ -2507,18 +2501,18 @@ lemma real_of_ereal_mult[simp]: fixes a b :: ereal - shows "real (a * b) = real a * real b" + shows "real_of_ereal (a * b) = real_of_ereal a * real_of_ereal b" by (cases rule: ereal2_cases[of a b]) auto lemma real_of_ereal_eq_0: fixes x :: ereal - shows "real x = 0 \ x = \ \ x = -\ \ x = 0" + shows "real_of_ereal x = 0 \ x = \ \ x = -\ \ x = 0" by (cases x) auto lemma tendsto_ereal_realD: fixes f :: "'a \ ereal" assumes "x \ 0" - and tendsto: "((\x. ereal (real (f x))) ---> x) net" + and tendsto: "((\x. ereal (real_of_ereal (f x))) ---> x) net" shows "(f ---> x) net" proof (intro topological_tendstoI) fix S @@ -2533,14 +2527,14 @@ lemma tendsto_ereal_realI: fixes f :: "'a \ ereal" assumes x: "\x\ \ \" and tendsto: "(f ---> x) net" - shows "((\x. ereal (real (f x))) ---> x) net" + shows "((\x. ereal (real_of_ereal (f x))) ---> x) net" proof (intro topological_tendstoI) fix S assume "open S" and "x \ S" with x have "open (S - {\, -\})" "x \ S - {\, -\}" by auto from tendsto[THEN topological_tendstoD, OF this] - show "eventually (\x. ereal (real (f x)) \ S) net" + show "eventually (\x. ereal (real_of_ereal (f x)) \ S) net" by (elim eventually_elim1) (auto simp: ereal_real) qed @@ -2556,15 +2550,15 @@ shows "((\x. f x + g x) ---> x + y) F" proof - from x obtain r where x': "x = ereal r" by (cases x) auto - with f have "((\i. real (f i)) ---> r) F" by simp + with f have "((\i. real_of_ereal (f i)) ---> r) F" by simp moreover from y obtain p where y': "y = ereal p" by (cases y) auto - with g have "((\i. real (g i)) ---> p) F" by simp - ultimately have "((\i. real (f i) + real (g i)) ---> r + p) F" + with g have "((\i. real_of_ereal (g i)) ---> p) F" by simp + ultimately have "((\i. real_of_ereal (f i) + real_of_ereal (g i)) ---> r + p) F" by (rule tendsto_add) moreover from eventually_finite[OF x f] eventually_finite[OF y g] - have "eventually (\x. f x + g x = ereal (real (f x) + real (g x))) F" + have "eventually (\x. f x + g x = ereal (real_of_ereal (f x) + real_of_ereal (g x))) F" by eventually_elim auto ultimately show ?thesis by (simp add: x' y' cong: filterlim_cong) @@ -2614,14 +2608,14 @@ lemma ereal_real': assumes "\x\ \ \" - shows "ereal (real x) = x" + shows "ereal (real_of_ereal x) = x" using assms by auto -lemma real_ereal_id: "real \ ereal = id" +lemma real_ereal_id: "real_of_ereal \ ereal = id" proof - { fix x - have "(real o ereal) x = id x" + have "(real_of_ereal o ereal) x = id x" by auto } then show ?thesis @@ -2682,7 +2676,7 @@ then have "rx < ra + r" and "ra < rx + r" using rx assms \0 < r\ lower[OF \n \ N\] upper[OF \n \ N\] by auto - then have "dist (real (u N)) rx < r" + then have "dist (real_of_ereal (u N)) rx < r" using rx ra_def by (auto simp: dist_real_def abs_diff_less_iff field_simps) from dist[OF this] show "u N \ S" @@ -3063,7 +3057,7 @@ fixes f :: "nat \ ereal" assumes f: "\i. 0 \ f i" and fin: "(\i. f i) \ \" - shows "summable (\i. real (f i))" + shows "summable (\i. real_of_ereal (f i))" proof (rule summable_def[THEN iffD2]) have "0 \ (\i. f i)" using assms by (auto intro: suminf_0_le) @@ -3077,12 +3071,12 @@ using f[of i] by auto } note fin = this - have "(\i. ereal (real (f i))) sums (\i. ereal (real (f i)))" + have "(\i. ereal (real_of_ereal (f i))) sums (\i. ereal (real_of_ereal (f i)))" using f by (auto intro!: summable_ereal_pos simp: ereal_le_real_iff zero_ereal_def) also have "\ = ereal r" using fin r by (auto simp: ereal_real) - finally show "\r. (\i. real (f i)) sums r" + finally show "\r. (\i. real_of_ereal (f i)) sums r" by (auto simp: sums_ereal) qed @@ -3559,7 +3553,7 @@ subsubsection \Continuity\ lemma continuous_at_of_ereal: - "\x0 :: ereal\ \ \ \ continuous (at x0) real" + "\x0 :: ereal\ \ \ \ continuous (at x0) real_of_ereal" unfolding continuous_at by (rule lim_real_of_ereal) (simp add: ereal_real) @@ -3583,10 +3577,10 @@ by (auto simp add: ereal_all_split ereal_ex_split) lemma ereal_tendsto_simps1: - "((f \ real) ---> y) (at_left (ereal x)) \ (f ---> y) (at_left x)" - "((f \ real) ---> y) (at_right (ereal x)) \ (f ---> y) (at_right x)" - "((f \ real) ---> y) (at_left (\::ereal)) \ (f ---> y) at_top" - "((f \ real) ---> y) (at_right (-\::ereal)) \ (f ---> y) at_bot" + "((f \ real_of_ereal) ---> y) (at_left (ereal x)) \ (f ---> y) (at_left x)" + "((f \ real_of_ereal) ---> y) (at_right (ereal x)) \ (f ---> y) (at_right x)" + "((f \ real_of_ereal) ---> y) (at_left (\::ereal)) \ (f ---> y) at_top" + "((f \ real_of_ereal) ---> y) (at_right (-\::ereal)) \ (f ---> y) at_bot" unfolding tendsto_compose_filtermap at_left_ereal at_right_ereal at_left_PInf at_right_MInf by (auto simp: filtermap_filtermap filtermap_ident) @@ -3638,24 +3632,24 @@ shows "continuous_on A f \ continuous_on A (ereal \ f)" unfolding continuous_on_def comp_def lim_ereal .. -lemma continuous_on_real: "continuous_on (UNIV - {\, -\::ereal}) real" +lemma continuous_on_real: "continuous_on (UNIV - {\, -\::ereal}) real_of_ereal" using continuous_at_of_ereal continuous_on_eq_continuous_at open_image_ereal by auto lemma continuous_on_iff_real: fixes f :: "'a::t2_space \ ereal" assumes *: "\x. x \ A \ \f x\ \ \" - shows "continuous_on A f \ continuous_on A (real \ f)" + shows "continuous_on A f \ continuous_on A (real_of_ereal \ f)" proof - have "f ` A \ UNIV - {\, -\}" using assms by force - then have *: "continuous_on (f ` A) real" + then have *: "continuous_on (f ` A) real_of_ereal" using continuous_on_real by (simp add: continuous_on_subset) - have **: "continuous_on ((real \ f) ` A) ereal" + have **: "continuous_on ((real_of_ereal \ f) ` A) ereal" by (intro continuous_on_ereal continuous_on_id) { assume "continuous_on A f" - then have "continuous_on A (real \ f)" + then have "continuous_on A (real_of_ereal \ f)" apply (subst continuous_on_compose) using * apply auto @@ -3663,14 +3657,14 @@ } moreover { - assume "continuous_on A (real \ f)" - then have "continuous_on A (ereal \ (real \ f))" + assume "continuous_on A (real_of_ereal \ f)" + then have "continuous_on A (ereal \ (real_of_ereal \ f))" apply (subst continuous_on_compose) using ** apply auto done then have "continuous_on A f" - apply (subst continuous_on_cong[of _ A _ "ereal \ (real \ f)"]) + apply (subst continuous_on_cong[of _ A _ "ereal \ (real_of_ereal \ f)"]) using assms ereal_real apply auto done @@ -3688,6 +3682,6 @@ value "\-\\ :: ereal" value "4 + 5 / 4 - ereal 2 :: ereal" value "ereal 3 < \" -value "real (\::ereal) = 0" +value "real_of_ereal (\::ereal) = 0" end diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Library/Float.thy Tue Nov 10 14:18:41 2015 +0000 @@ -15,35 +15,21 @@ morphisms real_of_float float_of unfolding float_def by auto -instantiation float :: real_of -begin - -definition real_float :: "float \ real" where - real_of_float_def[code_unfold]: "real \ real_of_float" +setup_lifting type_definition_float -instance .. - -end - -lemma type_definition_float': "type_definition real float_of float" - using type_definition_float unfolding real_of_float_def . - -setup_lifting type_definition_float' +declare real_of_float [code_unfold] lemmas float_of_inject[simp] -declare [[coercion "real :: float \ real"]] +declare [[coercion "real_of_float :: float \ real"]] lemma real_of_float_eq: fixes f1 f2 :: float - shows "f1 = f2 \ real f1 = real f2" - unfolding real_of_float_def real_of_float_inject .. + shows "f1 = f2 \ real_of_float f1 = real_of_float f2" + unfolding real_of_float_inject .. -lemma float_of_real[simp]: "float_of (real x) = x" - unfolding real_of_float_def by (rule real_of_float_inverse) - -lemma real_float[simp]: "x \ float \ real (float_of x) = x" - unfolding real_of_float_def by (rule float_of_inverse) +declare real_of_float_inverse[simp] float_of_inverse [simp] +declare real_of_float [simp] subsection \Real operations preserving the representation as floating point number\ @@ -59,15 +45,15 @@ by (intro floatI[of "numeral i" 0]) simp lemma neg_numeral_float[simp]: "- numeral i \ float" by (intro floatI[of "- numeral i" 0]) simp -lemma real_of_int_float[simp]: "real (x :: int) \ float" +lemma real_of_int_float[simp]: "real_of_int (x :: int) \ float" by (intro floatI[of x 0]) simp lemma real_of_nat_float[simp]: "real (x :: nat) \ float" by (intro floatI[of x 0]) simp -lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \ float" +lemma two_powr_int_float[simp]: "2 powr (real_of_int (i::int)) \ float" by (intro floatI[of 1 i]) simp lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \ float" by (intro floatI[of 1 i]) simp -lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \ float" +lemma two_powr_minus_int_float[simp]: "2 powr - (real_of_int (i::int)) \ float" by (intro floatI[of 1 "-i"]) simp lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \ float" by (intro floatI[of 1 "-i"]) simp @@ -77,8 +63,8 @@ by (intro floatI[of 1 "- numeral i"]) simp lemma two_pow_float[simp]: "2 ^ n \ float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow) -lemma real_of_float_float[simp]: "real (f::float) \ float" - by (cases f) simp + + lemma plus_float[simp]: "r \ float \ p \ float \ r + p \ float" unfolding float_def @@ -188,7 +174,7 @@ lemma compute_real_of_float[code]: "real_of_float (Float m e) = (if e \ 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))" - by (simp add: real_of_float_def[symmetric] powr_int) + by (simp add: powr_int) code_datatype Float @@ -233,13 +219,13 @@ lemma real_of_float_power[simp]: fixes f :: float - shows "real (f^n) = real f^n" + shows "real_of_float (f^n) = real_of_float f^n" by (induct n) simp_all lemma fixes x y :: float - shows real_of_float_min: "real (min x y) = min (real x) (real y)" - and real_of_float_max: "real (max x y) = max (real x) (real y)" + shows real_of_float_min: "real_of_float (min x y) = min (real_of_float x) (real_of_float y)" + and real_of_float_max: "real_of_float (max x y) = max (real_of_float x) (real_of_float y)" by (simp_all add: min_def max_def) instance float :: unbounded_dense_linorder @@ -277,10 +263,10 @@ end -lemma float_numeral[simp]: "real (numeral x :: float) = numeral x" +lemma float_numeral[simp]: "real_of_float (numeral x :: float) = numeral x" apply (induct x) apply simp - apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq real_float + apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq float_of_inverse plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float) done @@ -288,7 +274,7 @@ "rel_fun (op =) pcr_float (numeral :: _ \ real) (numeral :: _ \ float)" by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def) -lemma float_neg_numeral[simp]: "real (- numeral x :: float) = - numeral x" +lemma float_neg_numeral[simp]: "real_of_float (- numeral x :: float) = - numeral x" by simp lemma transfer_neg_numeral [transfer_rule]: @@ -384,7 +370,7 @@ also have "\ = m2 * 2^nat (e2 - e1)" by (simp add: powr_realpow) finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)" - unfolding real_of_int_inject . + by blast with m1 have "m1 = m2" by (cases "nat (e2 - e1)") (auto simp add: dvd_def) then show ?thesis @@ -422,7 +408,7 @@ lemma float_normed_cases: fixes f :: float obtains (zero) "f = 0" - | (powr) m e :: int where "real f = m * 2 powr e" "\ 2 dvd m" "f \ 0" + | (powr) m e :: int where "real_of_float f = m * 2 powr e" "\ 2 dvd m" "f \ 0" proof (atomize_elim, induct f) case (float_of y) then show ?case @@ -431,11 +417,11 @@ definition mantissa :: "float \ int" where "mantissa f = fst (SOME p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) - \ (f \ 0 \ real f = real (fst p) * 2 powr real (snd p) \ \ 2 dvd fst p))" + \ (f \ 0 \ real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \ \ 2 dvd fst p))" definition exponent :: "float \ int" where "exponent f = snd (SOME p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) - \ (f \ 0 \ real f = real (fst p) * 2 powr real (snd p) \ \ 2 dvd fst p))" + \ (f \ 0 \ real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \ \ 2 dvd fst p))" lemma shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E) @@ -448,7 +434,7 @@ qed lemma - shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E) + shows mantissa_exponent: "real_of_float f = mantissa f * 2 powr exponent f" (is ?E) and mantissa_not_dvd: "f \ (float_of 0) \ \ 2 dvd mantissa f" (is "_ \ ?D") proof cases assume [simp]: "f \ float_of 0" @@ -459,7 +445,7 @@ next case (powr m e) then have "\p::int \ int. (f = 0 \ fst p = 0 \ snd p = 0) \ - (f \ 0 \ real f = real (fst p) * 2 powr real (snd p) \ \ 2 dvd fst p)" + (f \ 0 \ real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \ \ 2 dvd fst p)" by auto then show ?thesis unfolding exponent_def mantissa_def @@ -517,14 +503,14 @@ proof (rule ccontr) assume "\ e \ exponent f" then have pos: "exponent f < e" by simp - then have "2 powr (exponent f - e) = 2 powr - real (e - exponent f)" + then have "2 powr (exponent f - e) = 2 powr - real_of_int (e - exponent f)" by simp also have "\ = 1 / 2^nat (e - exponent f)" using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric]) - finally have "m * 2^nat (e - exponent f) = real (mantissa f)" + finally have "m * 2^nat (e - exponent f) = real_of_int (mantissa f)" using eq by simp then have "mantissa f = m * 2^nat (e - exponent f)" - unfolding real_of_int_inject by simp + by linarith with \exponent f < e\ have "2 dvd mantissa f" apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"]) apply (cases "nat (e - exponent f)") @@ -532,11 +518,11 @@ done then show False using mantissa_not_dvd[OF not_0] by simp qed - ultimately have "real m = mantissa f * 2^nat (exponent f - e)" + ultimately have "real_of_int m = mantissa f * 2^nat (exponent f - e)" by (simp add: powr_realpow[symmetric]) with \e \ exponent f\ show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)" - unfolding real_of_int_inject by auto + by force+ qed context @@ -613,25 +599,14 @@ subsection \Lemmas for types @{typ real}, @{typ nat}, @{typ int}\ lemmas real_of_ints = - real_of_int_zero - real_of_one - real_of_int_add - real_of_int_minus - real_of_int_diff - real_of_int_mult - real_of_int_power - real_numeral -lemmas real_of_nats = - real_of_nat_zero - real_of_nat_one - real_of_nat_1 - real_of_nat_add - real_of_nat_mult - real_of_nat_power - real_of_nat_numeral + of_int_add + of_int_minus + of_int_diff + of_int_mult + of_int_power + of_int_numeral of_int_neg_numeral lemmas int_of_reals = real_of_ints[symmetric] -lemmas nat_of_reals = real_of_nats[symmetric] subsection \Rounding Real Numbers\ @@ -644,14 +619,14 @@ lemma round_down_float[simp]: "round_down prec x \ float" unfolding round_down_def - by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus) + by (auto intro!: times_float simp: of_int_minus[symmetric] simp del: of_int_minus) lemma round_up_float[simp]: "round_up prec x \ float" unfolding round_up_def - by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus) + by (auto intro!: times_float simp: of_int_minus[symmetric] simp del: of_int_minus) lemma round_up: "x \ round_up prec x" - by (simp add: powr_minus_divide le_divide_eq round_up_def) + by (simp add: powr_minus_divide le_divide_eq round_up_def ceiling_correct) lemma round_down: "round_down prec x \ x" by (simp add: powr_minus_divide divide_le_eq round_down_def) @@ -670,8 +645,8 @@ by (simp add: round_up_def round_down_def field_simps) also have "\ \ 1 * 2 powr -prec" by (rule mult_mono) - (auto simp del: real_of_int_diff - simp: real_of_int_diff[symmetric] real_of_int_le_one_cancel_iff ceiling_diff_floor_le_1) + (auto simp del: of_int_diff + simp: of_int_diff[symmetric] ceiling_diff_floor_le_1) finally show ?thesis by simp qed @@ -696,7 +671,7 @@ assumes "x \ 1" "prec \ 0" shows "round_up prec x \ 1" proof - - have "real \x * 2 powr prec\ \ real \2 powr real prec\" + have "real_of_int \x * 2 powr prec\ \ real_of_int \2 powr real_of_int prec\" using assms by (auto intro!: ceiling_mono) also have "\ = 2 powr prec" using assms by (auto simp: powr_int intro!: exI[where x="2^nat prec"]) finally show ?thesis @@ -712,7 +687,7 @@ also have "\ \ 2 powr p - 1" using \p > 0\ by (auto simp: powr_divide2[symmetric] powr_int field_simps self_le_power) finally show ?thesis using \p > 0\ - by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_eq) + by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_iff) qed lemma round_down_ge1: @@ -721,9 +696,9 @@ shows "1 \ round_down p x" proof cases assume nonneg: "0 \ p" - have "2 powr p = real \2 powr real p\" + have "2 powr p = real_of_int \2 powr real_of_int p\" using nonneg by (auto simp: powr_int) - also have "\ \ real \x * 2 powr p\" + also have "\ \ real_of_int \x * 2 powr p\" using assms by (auto intro!: floor_mono) finally show ?thesis by (simp add: round_down_def) (simp add: powr_minus inverse_eq_divide) @@ -735,11 +710,11 @@ using prec by auto finally have x_le: "x \ 2 powr -p" . - from neg have "2 powr real p \ 2 powr 0" + from neg have "2 powr real_of_int p \ 2 powr 0" by (intro powr_mono) auto also have "\ \ \2 powr 0::real\" by simp - also have "\ \ \x * 2 powr (real p)\" - unfolding real_of_int_le_iff + also have "\ \ \x * 2 powr (real_of_int p)\" + unfolding of_int_le_iff using x x_le by (intro floor_mono) (simp add: powr_minus_divide field_simps) finally show ?thesis using prec x @@ -777,11 +752,11 @@ using round_down by simp also have "\ \ 2 powr -e" using round_up_diff_round_down by simp - finally show "round_up e f - f \ 2 powr - (real e)" + finally show "round_up e f - f \ 2 powr - (real_of_int e)" by simp qed (simp add: algebra_simps round_up) -lemma float_up_correct: "real (float_up e f) - real f \ {0..2 powr -e}" +lemma float_up_correct: "real_of_float (float_up e f) - real_of_float f \ {0..2 powr -e}" by transfer (rule round_up_correct) lift_definition float_down :: "int \ float \ float" is round_down by simp @@ -794,11 +769,11 @@ using round_up by simp also have "\ \ 2 powr -e" using round_up_diff_round_down by simp - finally show "f - round_down e f \ 2 powr - (real e)" + finally show "f - round_down e f \ 2 powr - (real_of_int e)" by simp qed (simp add: algebra_simps round_down) -lemma float_down_correct: "real f - real (float_down e f) \ {0..2 powr -e}" +lemma float_down_correct: "real_of_float f - real_of_float (float_down e f) \ {0..2 powr -e}" by transfer (rule round_down_correct) context @@ -809,17 +784,27 @@ (if p + e < 0 then Float (div_twopow m (nat (-(p + e)))) (-p) else Float m e)" proof (cases "p + e < 0") case True - then have "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" + then have "real_of_int ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))" using powr_realpow[of 2 "nat (-(p + e))"] by simp also have "\ = 1 / 2 powr p / 2 powr e" - unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add) + unfolding powr_minus_divide of_int_minus by (simp add: powr_add) finally show ?thesis using \p + e < 0\ - by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric]) + apply transfer + apply (simp add: ac_simps round_down_def floor_divide_of_int_eq[symmetric]) + proof - (*FIXME*) + fix pa :: int and ea :: int and ma :: int + assume a1: "2 ^ nat (- pa - ea) = 1 / (2 powr real_of_int pa * 2 powr real_of_int ea)" + assume "pa + ea < 0" + have "\real_of_int ma / real_of_int (int 2 ^ nat (- (pa + ea)))\ = \real_of_float (Float ma (pa + ea))\" + using a1 by (simp add: powr_add) + thus "\real_of_int ma * (2 powr real_of_int pa * 2 powr real_of_int ea)\ = ma div 2 ^ nat (- pa - ea)" + by (metis Float.rep_eq add_uminus_conv_diff floor_divide_of_int_eq minus_add_distrib of_int_simps(3) of_nat_numeral powr_add) + qed next case False - then have r: "real e + real p = real (nat (e + p))" by simp - have r: "\(m * 2 powr e) * 2 powr real p\ = (m * 2 powr e) * 2 powr real p" + then have r: "real_of_int e + real_of_int p = real (nat (e + p))" by simp + have r: "\(m * 2 powr e) * 2 powr real_of_int p\ = (m * 2 powr e) * 2 powr real_of_int p" by (auto intro: exI[where x="m*2^nat (e+p)"] simp add: ac_simps powr_add[symmetric] r powr_realpow) with \\ p + e < 0\ show ?thesis @@ -837,30 +822,30 @@ lemma ceil_divide_floor_conv: assumes "b \ 0" - shows "\real a / real b\ = (if b dvd a then a div b else \real a / real b\ + 1)" + shows "\real_of_int a / real_of_int b\ = (if b dvd a then a div b else \real_of_int a / real_of_int b\ + 1)" proof (cases "b dvd a") case True then show ?thesis - by (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric] - floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus) + by (simp add: ceiling_def of_int_minus[symmetric] divide_minus_left[symmetric] + floor_divide_of_int_eq dvd_neg_div del: divide_minus_left of_int_minus) next case False then have "a mod b \ 0" by auto - then have ne: "real (a mod b) / real b \ 0" + then have ne: "real_of_int (a mod b) / real_of_int b \ 0" using \b \ 0\ by auto - have "\real a / real b\ = \real a / real b\ + 1" + have "\real_of_int a / real_of_int b\ = \real_of_int a / real_of_int b\ + 1" apply (rule ceiling_eq) - apply (auto simp: floor_divide_eq_div[symmetric]) + apply (auto simp: floor_divide_of_int_eq[symmetric]) proof - - have "real \real a / real b\ \ real a / real b" + have "real_of_int \real_of_int a / real_of_int b\ \ real_of_int a / real_of_int b" by simp - moreover have "real \real a / real b\ \ real a / real b" + moreover have "real_of_int \real_of_int a / real_of_int b\ \ real_of_int a / real_of_int b" apply (subst (2) real_of_int_div_aux) - unfolding floor_divide_eq_div + unfolding floor_divide_of_int_eq using ne \b \ 0\ apply auto done - ultimately show "real \real a / real b\ < real a / real b" by arith + ultimately show "real_of_int \real_of_int a / real_of_int b\ < real_of_int a / real_of_int b" by arith qed then show ?thesis using \\ b dvd a\ by simp @@ -897,15 +882,14 @@ proof show "2 ^ nat (bitlen x - 1) \ x" proof - - have "(2::real) ^ nat \log 2 (real x)\ = 2 powr real (floor (log 2 (real x)))" - using powr_realpow[symmetric, of 2 "nat \log 2 (real x)\"] \x > 0\ - using real_nat_eq_real[of "floor (log 2 (real x))"] + have "(2::real) ^ nat \log 2 (real_of_int x)\ = 2 powr real_of_int (floor (log 2 (real_of_int x)))" + using powr_realpow[symmetric, of 2 "nat \log 2 (real_of_int x)\"] \x > 0\ by simp - also have "\ \ 2 powr log 2 (real x)" + also have "\ \ 2 powr log 2 (real_of_int x)" by simp - also have "\ = real x" + also have "\ = real_of_int x" using \0 < x\ by simp - finally have "2 ^ nat \log 2 (real x)\ \ real x" + finally have "2 ^ nat \log 2 (real_of_int x)\ \ real_of_int x" by simp then show ?thesis using \0 < x\ by (simp add: bitlen_def) @@ -914,7 +898,7 @@ proof - have "x \ 2 powr (log 2 x)" using \x > 0\ by simp - also have "\ < 2 ^ nat (\log 2 (real x)\ + 1)" + also have "\ < 2 ^ nat (\log 2 (real_of_int x)\ + 1)" apply (simp add: powr_realpow[symmetric]) using \x > 0\ apply simp done @@ -930,8 +914,9 @@ from assms have "b * 2 ^ c > 0" by auto then show ?thesis - using floor_add[of "log 2 b" c] assms - by (auto simp add: log_mult log_nat_power bitlen_def) + using floor_add[of "log 2 b" c] assms + apply (auto simp add: log_mult log_nat_power bitlen_def) + by (metis add.right_neutral frac_lt_1 frac_of_int of_int_of_nat_eq) qed lemma bitlen_Float: @@ -961,13 +946,13 @@ { assume "2 \ x" then have "\log 2 (x div 2)\ + 1 = \log 2 (x - x mod 2)\" by (simp add: log_mult zmod_zdiv_equality') - also have "\ = \log 2 (real x)\" + also have "\ = \log 2 (real_of_int x)\" proof (cases "x mod 2 = 0") case True then show ?thesis by simp next case False - def n \ "\log 2 (real x)\" + def n \ "\log 2 (real_of_int x)\" then have "0 \ n" using \2 \ x\ by simp from \2 \ x\ False have "x mod 2 = 1" "\ 2 dvd x" @@ -975,18 +960,18 @@ with \2 \ x\ have "x \ 2 ^ nat n" by (cases "nat n") auto moreover - { have "real (2^nat n :: int) = 2 powr (nat n)" + { have "real_of_int (2^nat n :: int) = 2 powr (nat n)" by (simp add: powr_realpow) also have "\ \ 2 powr (log 2 x)" using \2 \ x\ by (simp add: n_def del: powr_log_cancel) finally have "2^nat n \ x" using \2 \ x\ by simp } ultimately have "2^nat n \ x - 1" by simp - then have "2^nat n \ real (x - 1)" - unfolding real_of_int_le_iff[symmetric] by simp + then have "2^nat n \ real_of_int (x - 1)" + using numeral_power_le_real_of_int_cancel_iff by blast { have "n = \log 2 (2^nat n)\" using \0 \ n\ by (simp add: log_nat_power) also have "\ \ \log 2 (x - 1)\" - using \2^nat n \ real (x - 1)\ \0 \ n\ \2 \ x\ by (auto intro: floor_mono) + using \2^nat n \ real_of_int (x - 1)\ \0 \ n\ \2 \ x\ by (auto intro: floor_mono) finally have "n \ \log 2 (x - 1)\" . } moreover have "\log 2 (x - 1)\ \ n" using \2 \ x\ by (auto simp add: n_def intro!: floor_mono) @@ -997,7 +982,7 @@ moreover { assume "x < 2" "0 < x" then have "x = 1" by simp - then have "\log 2 (real x)\ = 0" by simp } + then have "\log 2 (real_of_int x)\ = 0" by simp } ultimately show ?thesis unfolding bitlen_def by (auto simp: pos_imp_zdiv_pos_iff not_le) @@ -1026,15 +1011,15 @@ have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"] by (auto simp: powr_minus field_simps) - then have "1 \ real m * inverse ?S" + then have "1 \ real_of_int m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"] by (auto simp: powr_minus) - then have "1 * ?S \ real m * inverse ?S * ?S" + then have "1 * ?S \ real_of_int m * inverse ?S * ?S" by (rule mult_right_mono) auto - then have "?S \ real m" + then have "?S \ real_of_int m" unfolding mult.assoc by auto then have "?S \ m" - unfolding real_of_int_le_iff[symmetric] by auto + unfolding of_int_le_iff[symmetric] by auto from this bitlen_bounds[OF \0 < m\, THEN conjunct2] have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF \1 < 2\, symmetric] @@ -1048,15 +1033,15 @@ lemma bitlen_div: assumes "0 < m" - shows "1 \ real m / 2^nat (bitlen m - 1)" - and "real m / 2^nat (bitlen m - 1) < 2" + shows "1 \ real_of_int m / 2^nat (bitlen m - 1)" + and "real_of_int m / 2^nat (bitlen m - 1) < 2" proof - let ?B = "2^nat(bitlen m - 1)" have "?B \ m" using bitlen_bounds[OF \0 ] .. - then have "1 * ?B \ real m" - unfolding real_of_int_le_iff[symmetric] by auto - then show "1 \ real m / ?B" + then have "1 * ?B \ real_of_int m" + unfolding of_int_le_iff[symmetric] by auto + then show "1 \ real_of_int m / ?B" by auto have "m \ 0" @@ -1070,11 +1055,11 @@ using \0 < m\ by (auto simp: bitlen_def) also have "\ = ?B * 2" unfolding nat_add_distrib[OF \0 \ bitlen m - 1\ zero_le_one] by auto - finally have "real m < 2 * ?B" - unfolding real_of_int_less_iff[symmetric] by auto - then have "real m / ?B < 2 * ?B / ?B" + finally have "real_of_int m < 2 * ?B" + by (metis (full_types) mult.commute power.simps(2) real_of_int_less_numeral_power_cancel_iff) + then have "real_of_int m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono) auto - then show "real m / ?B < 2" + then show "real_of_int m / ?B < 2" by auto qed @@ -1122,7 +1107,7 @@ assumes "x > 0" "p > 0" shows "truncate_down p x > 0" proof - - have "0 \ log 2 x - real \log 2 x\" + have "0 \ log 2 x - real_of_int \log 2 x\" by (simp add: algebra_simps) from this assms show ?thesis @@ -1167,7 +1152,7 @@ lift_definition float_round_up :: "nat \ float \ float" is truncate_up by (simp add: truncate_up_def) -lemma float_round_up: "real x \ real (float_round_up prec x)" +lemma float_round_up: "real_of_float x \ real_of_float (float_round_up prec x)" using truncate_up by transfer simp lemma float_round_up_zero[simp]: "float_round_up prec 0 = 0" @@ -1176,7 +1161,7 @@ lift_definition float_round_down :: "nat \ float \ float" is truncate_down by (simp add: truncate_down_def) -lemma float_round_down: "real (float_round_down prec x) \ real x" +lemma float_round_down: "real_of_float (float_round_down prec x) \ real_of_float x" using truncate_down by transfer simp lemma float_round_down_zero[simp]: "float_round_down prec 0 = 0" @@ -1216,8 +1201,8 @@ lemma real_div_nat_eq_floor_of_divide: fixes a b :: nat - shows "a div b = real \a / b\" - by (simp add: floor_divide_of_nat_eq [of a b] real_eq_of_nat) + shows "a div b = real_of_int \a / b\" + by (simp add: floor_divide_of_nat_eq [of a b]) definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)" @@ -1269,25 +1254,28 @@ def x' \ "x * 2 ^ nat l" have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power) - moreover have "real x * 2 powr real l = real x'" + moreover have "real x * 2 powr l = real x'" by (simp add: powr_realpow[symmetric] \0 \ l\ x'_def) ultimately show ?thesis using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] \0 \ l\ \y \ 0\ l_def[symmetric, THEN meta_eq_to_obj_eq] - by transfer (auto simp add: floor_divide_eq_div [symmetric] round_up_def) + apply transfer + apply (auto simp add: round_up_def) + by (metis floor_divide_of_int_eq of_int_of_nat_eq) next case False def y' \ "y * 2 ^ nat (- l)" from \y \ 0\ have "y' \ 0" by (simp add: y'_def) have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power) - moreover have "real x * real (2::int) powr real l / real y = x / real y'" + moreover have "real x * real_of_int (2::int) powr real_of_int l / real y = x / real y'" using \\ 0 \ l\ by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps) ultimately show ?thesis using ceil_divide_floor_conv[of y' x] \\ 0 \ l\ \y' \ 0\ \y \ 0\ l_def[symmetric, THEN meta_eq_to_obj_eq] - by transfer - (auto simp add: round_up_def ceil_divide_floor_conv floor_divide_eq_div [symmetric]) + apply transfer + apply (auto simp add: round_up_def ceil_divide_floor_conv) + by (metis floor_divide_of_int_eq of_int_of_nat_eq) qed qed @@ -1312,7 +1300,7 @@ qed lemma rapprox_posrat_less1: - "0 \ x \ 0 < y \ 2 * x < y \ 0 < n \ real (rapprox_posrat n x y) < 1" + "0 \ x \ 0 < y \ 2 * x < y \ 0 < n \ real_of_float (rapprox_posrat n x y) < 1" by transfer (simp add: rat_precision_pos round_up_less1) lift_definition lapprox_rat :: "nat \ int \ int \ float" is @@ -1366,12 +1354,12 @@ "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)" proof (cases "m1 \ 0 \ m2 \ 0") case True - let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2" - let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)" + let ?f1 = "real_of_int m1 * 2 powr real_of_int s1" and ?f2 = "real_of_int m2 * 2 powr real_of_int s2" + let ?m = "real_of_int m1 / real_of_int m2" and ?s = "2 powr real_of_int (s1 - s2)" from True have eq2: "(int prec + \log 2 \?f2\\ - \log 2 \?f1\\) = rat_precision prec \m1\ \m2\ + (s2 - s1)" by (simp add: abs_mult log_mult rat_precision_def bitlen_def) - have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s" + have eq1: "real_of_int m1 * 2 powr real_of_int s1 / (real_of_int m2 * 2 powr real_of_int s2) = ?m * ?s" by (simp add: field_simps powr_divide2[symmetric]) from True show ?thesis by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift real_divl_def, @@ -1494,10 +1482,10 @@ lemma power_up_fl: "0 \ x \ x ^ n \ power_up_fl p x n" by transfer (rule power_up) -lemma real_power_up_fl: "real (power_up_fl p x n) = power_up p x n" +lemma real_power_up_fl: "real_of_float (power_up_fl p x n) = power_up p x n" by transfer simp -lemma real_power_down_fl: "real (power_down_fl p x n) = power_down p x n" +lemma real_power_down_fl: "real_of_float (power_down_fl p x n) = power_down p x n" by transfer simp @@ -1521,8 +1509,8 @@ and plus_up: "x + y \ plus_up prec x y" by (auto simp: plus_down_def truncate_down plus_up_def truncate_up) -lemma float_plus_down: "real (float_plus_down prec x y) \ x + y" - and float_plus_up: "x + y \ real (float_plus_up prec x y)" +lemma float_plus_down: "real_of_float (float_plus_down prec x y) \ x + y" + and float_plus_up: "x + y \ real_of_float (float_plus_up prec x y)" by (transfer, rule plus_down plus_up)+ lemmas plus_down_le = order_trans[OF plus_down] @@ -1551,7 +1539,7 @@ and "abs a > k \ abs b \ k \ a + b \ 0" by auto -lemma abs_real_le_2_powr_bitlen[simp]: "\real m2\ < 2 powr real (bitlen \m2\)" +lemma abs_real_le_2_powr_bitlen[simp]: "\real_of_int m2\ < 2 powr real_of_int (bitlen \m2\)" proof (cases "m2 = 0") case True then show ?thesis by simp @@ -1561,7 +1549,7 @@ using bitlen_bounds[of "\m2\"] by (auto simp: powr_add bitlen_nonneg) then show ?thesis - by (simp add: powr_int bitlen_nonneg real_of_int_less_iff[symmetric]) + by (metis bitlen_nonneg powr_int real_of_int_abs real_of_int_less_numeral_power_cancel_iff zero_less_numeral) qed lemma floor_sum_times_2_powr_sgn_eq: @@ -1583,37 +1571,37 @@ then have "b * 2 powr p < abs (b * 2 powr (p + 1))" by simp also note b_le_1 - finally have b_less_1: "b * 2 powr real p < 1" . + finally have b_less_1: "b * 2 powr real_of_int p < 1" . - from b_less_1 \b > 0\ have floor_eq: "\b * 2 powr real p\ = 0" "\sgn b / 2\ = 0" + from b_less_1 \b > 0\ have floor_eq: "\b * 2 powr real_of_int p\ = 0" "\sgn b / 2\ = 0" by (simp_all add: floor_eq_iff) have "\(a + b) * 2 powr q\ = \(a + b) * 2 powr p * 2 powr (q - p)\" by (simp add: algebra_simps powr_realpow[symmetric] powr_add[symmetric]) also have "\ = \(ai + b * 2 powr p) * 2 powr (q - p)\" by (simp add: assms algebra_simps) - also have "\ = \(ai + b * 2 powr p) / real ((2::int) ^ nat (p - q))\" + also have "\ = \(ai + b * 2 powr p) / real_of_int ((2::int) ^ nat (p - q))\" using assms by (simp add: algebra_simps powr_realpow[symmetric] divide_powr_uminus powr_add[symmetric]) - also have "\ = \ai / real ((2::int) ^ nat (p - q))\" - by (simp del: real_of_int_power add: floor_divide_real_eq_div floor_eq) - finally have "\(a + b) * 2 powr real q\ = \real ai / real ((2::int) ^ nat (p - q))\" . + also have "\ = \ai / real_of_int ((2::int) ^ nat (p - q))\" + by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq) + finally have "\(a + b) * 2 powr real_of_int q\ = \real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\" . moreover { - have "\(2 * ai + sgn b) * 2 powr (real (q - p) - 1)\ = \(ai + sgn b / 2) * 2 powr (q - p)\" + have "\(2 * ai + sgn b) * 2 powr (real_of_int (q - p) - 1)\ = \(ai + sgn b / 2) * 2 powr (q - p)\" by (subst powr_divide2[symmetric]) (simp add: field_simps) - also have "\ = \(ai + sgn b / 2) / real ((2::int) ^ nat (p - q))\" + also have "\ = \(ai + sgn b / 2) / real_of_int ((2::int) ^ nat (p - q))\" using leqp by (simp add: powr_realpow[symmetric] powr_divide2[symmetric]) - also have "\ = \ai / real ((2::int) ^ nat (p - q))\" - by (simp del: real_of_int_power add: floor_divide_real_eq_div floor_eq) + also have "\ = \ai / real_of_int ((2::int) ^ nat (p - q))\" + by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq) finally - have "\(2 * ai + (sgn b)) * 2 powr (real (q - p) - 1)\ = - \real ai / real ((2::int) ^ nat (p - q))\" . + have "\(2 * ai + (sgn b)) * 2 powr (real_of_int (q - p) - 1)\ = + \real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\" . } ultimately show ?thesis by simp next case 3 - then have floor_eq: "\b * 2 powr (real p + 1)\ = -1" + then have floor_eq: "\b * 2 powr (real_of_int p + 1)\ = -1" using b_le_1 by (auto simp: floor_eq_iff algebra_simps pos_divide_le_eq[symmetric] abs_if divide_powr_uminus intro!: mult_neg_pos split: split_if_asm) @@ -1623,12 +1611,12 @@ by (simp add: algebra_simps) also have "\ = \(2 * ai + b * 2 powr (p + 1)) / 2 powr (1 - q + p)\" using assms by (simp add: algebra_simps powr_mult_base divide_powr_uminus) - also have "\ = \(2 * ai + b * 2 powr (p + 1)) / real ((2::int) ^ nat (p - q + 1))\" + also have "\ = \(2 * ai + b * 2 powr (p + 1)) / real_of_int ((2::int) ^ nat (p - q + 1))\" using assms by (simp add: algebra_simps powr_realpow[symmetric]) - also have "\ = \(2 * ai - 1) / real ((2::int) ^ nat (p - q + 1))\" + also have "\ = \(2 * ai - 1) / real_of_int ((2::int) ^ nat (p - q + 1))\" using \b < 0\ assms - by (simp add: floor_divide_eq_div floor_eq floor_divide_real_eq_div - del: real_of_int_mult real_of_int_power real_of_int_diff) + by (simp add: floor_divide_of_int_eq floor_eq floor_divide_real_eq_div + del: of_int_mult of_int_power of_int_diff) also have "\ = \(2 * ai - 1) * 2 powr (q - p - 1)\" using assms by (simp add: algebra_simps divide_powr_uminus powr_realpow[symmetric]) finally show ?thesis @@ -1641,7 +1629,7 @@ and b :: real assumes "abs b \ 1/2" and "ai \ 0" - shows "\log 2 \real ai + b\\ = \log 2 \ai + sgn b / 2\\" + shows "\log 2 \real_of_int ai + b\\ = \log 2 \ai + sgn b / 2\\" proof (cases "b = 0") case True then show ?thesis by simp @@ -1660,9 +1648,10 @@ by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int) then have "r \ (2::int) ^ nat k - 1" using \k \ 0\ by (auto simp: powr_int) - from this[simplified real_of_int_le_iff[symmetric]] \0 \ k\ + from this[simplified of_int_le_iff[symmetric]] \0 \ k\ have r_le: "r \ 2 powr k - 1" - by (auto simp: algebra_simps powr_int simp del: real_of_int_le_iff) + apply (auto simp: algebra_simps powr_int) + by (metis of_int_1 of_int_add real_of_int_le_numeral_power_cancel_iff) have "\ai\ = 2 powr k + r" using \k \ 0\ by (auto simp: k_def r_def powr_realpow[symmetric]) @@ -1680,7 +1669,7 @@ using \k \ 0\ r r_le by (auto simp: floor_log_eq_powr_iff powr_minus_divide field_simps sgn_if) - from \real \ai\ = _\ have "\ai + b\ = 2 powr k + (r + sgn ai * b)" + from \real_of_int \ai\ = _\ have "\ai + b\ = 2 powr k + (r + sgn ai * b)" using \abs b <= _\ \0 \ k\ r by (auto simp add: sgn_if abs_if) also have "\log 2 \\ = \log 2 (2 powr k + r + sgn (sgn ai * b) / 2)\" @@ -1708,7 +1697,7 @@ finally show ?thesis . qed also have "2 powr k + r + sgn (sgn ai * b) / 2 = \ai + sgn b / 2\" - unfolding \real \ai\ = _\[symmetric] using \ai \ 0\ + unfolding \real_of_int \ai\ = _\[symmetric] using \ai \ 0\ by (auto simp: abs_if sgn_if algebra_simps) finally show ?thesis . qed @@ -1724,10 +1713,10 @@ shows "float_plus_down p (Float m1 e1) (Float m2 e2) = float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2))" proof - - let ?a = "real (Float m1 e1)" - let ?b = "real (Float m2 e2)" + let ?a = "real_of_float (Float m1 e1)" + let ?b = "real_of_float (Float m2 e2)" let ?sum = "?a + ?b" - let ?shift = "real e2 - real e1 + real k1 + 1" + let ?shift = "real_of_int e2 - real_of_int e1 + real k1 + 1" let ?m1 = "m1 * 2 ^ Suc k1" let ?m2 = "m2 * 2 powr ?shift" let ?m2' = "sgn m2 / 2" @@ -1744,13 +1733,13 @@ finally have abs_m2_less_half: "\?m2\ < 1 / 2" by simp - then have "\real m2\ < 2 powr -(?shift + 1)" + then have "\real_of_int m2\ < 2 powr -(?shift + 1)" unfolding powr_minus_divide by (auto simp: bitlen_def field_simps powr_mult_base abs_mult) - also have "\ \ 2 powr real (e1 - e2 - 2)" + also have "\ \ 2 powr real_of_int (e1 - e2 - 2)" by simp - finally have b_less_quarter: "\?b\ < 1/4 * 2 powr real e1" + finally have b_less_quarter: "\?b\ < 1/4 * 2 powr real_of_int e1" by (simp add: powr_add field_simps powr_divide2[symmetric] powr_numeral abs_mult) - also have "1/4 < \real m1\ / 2" using \m1 \ 0\ by simp + also have "1/4 < \real_of_int m1\ / 2" using \m1 \ 0\ by simp finally have b_less_half_a: "\?b\ < 1/2 * \?a\" by (simp add: algebra_simps powr_mult_base abs_mult) then have a_half_less_sum: "\?a\ / 2 < \?sum\" @@ -1759,7 +1748,7 @@ from b_less_half_a have "\?b\ < \?a\" "\?b\ \ \?a\" by simp_all - have "\real (Float m1 e1)\ \ 1/4 * 2 powr real e1" + have "\real_of_float (Float m1 e1)\ \ 1/4 * 2 powr real_of_int e1" using \m1 \ 0\ by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult) then have "?sum \ 0" using b_less_quarter @@ -1767,51 +1756,50 @@ then have "?m1 + ?m2 \ 0" unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff) - have "\real ?m1\ \ 2 ^ Suc k1" "\?m2'\ < 2 ^ Suc k1" + have "\real_of_int ?m1\ \ 2 ^ Suc k1" "\?m2'\ < 2 ^ Suc k1" using \m1 \ 0\ \m2 \ 0\ by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps) then have sum'_nz: "?m1 + ?m2' \ 0" by (intro sum_neq_zeroI) - have "\log 2 \real (Float m1 e1) + real (Float m2 e2)\\ = \log 2 \?m1 + ?m2\\ + ?e" + have "\log 2 \real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\\ = \log 2 \?m1 + ?m2\\ + ?e" using \?m1 + ?m2 \ 0\ unfolding floor_add[symmetric] sum_eq - by (simp add: abs_mult log_mult) - also have "\log 2 \?m1 + ?m2\\ = \log 2 \?m1 + sgn (real m2 * 2 powr ?shift) / 2\\" + by (simp add: abs_mult log_mult) linarith + also have "\log 2 \?m1 + ?m2\\ = \log 2 \?m1 + sgn (real_of_int m2 * 2 powr ?shift) / 2\\" using abs_m2_less_half \m1 \ 0\ by (intro log2_abs_int_add_less_half_sgn_eq) (auto simp: abs_mult) - also have "sgn (real m2 * 2 powr ?shift) = sgn m2" + also have "sgn (real_of_int m2 * 2 powr ?shift) = sgn m2" by (auto simp: sgn_if zero_less_mult_iff less_not_sym) also have "\?m1 + ?m2'\ * 2 powr ?e = \?m1 * 2 + sgn m2\ * 2 powr (?e - 1)" by (auto simp: field_simps powr_minus[symmetric] powr_divide2[symmetric] powr_mult_base) - then have "\log 2 \?m1 + ?m2'\\ + ?e = \log 2 \real (Float (?m1 * 2 + sgn m2) (?e - 1))\\" + then have "\log 2 \?m1 + ?m2'\\ + ?e = \log 2 \real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1))\\" using \?m1 + ?m2' \ 0\ - unfolding floor_add[symmetric] + unfolding floor_add_of_int[symmetric] by (simp add: log_add_eq_powr abs_mult_pos) finally - have "\log 2 \?sum\\ = \log 2 \real (Float (?m1*2 + sgn m2) (?e - 1))\\" . + have "\log 2 \?sum\\ = \log 2 \real_of_float (Float (?m1*2 + sgn m2) (?e - 1))\\" . then have "plus_down p (Float m1 e1) (Float m2 e2) = truncate_down p (Float (?m1*2 + sgn m2) (?e - 1))" unfolding plus_down_def proof (rule truncate_down_log2_eqI) - let ?f = "(int p - \log 2 \real (Float m1 e1) + real (Float m2 e2)\\ - 1)" + let ?f = "(int p - \log 2 \real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\\ - 1)" let ?ai = "m1 * 2 ^ (Suc k1)" - have "\(?a + ?b) * 2 powr real ?f\ = \(real (2 * ?ai) + sgn ?b) * 2 powr real (?f - - ?e - 1)\" + have "\(?a + ?b) * 2 powr real_of_int ?f\ = \(real_of_int (2 * ?ai) + sgn ?b) * 2 powr real_of_int (?f - - ?e - 1)\" proof (rule floor_sum_times_2_powr_sgn_eq) - show "?a * 2 powr real (-?e) = real ?ai" + show "?a * 2 powr real_of_int (-?e) = real_of_int ?ai" by (simp add: powr_add powr_realpow[symmetric] powr_divide2[symmetric]) - show "\?b * 2 powr real (-?e + 1)\ \ 1" + show "\?b * 2 powr real_of_int (-?e + 1)\ \ 1" using abs_m2_less_half by (simp add: abs_mult powr_add[symmetric] algebra_simps powr_mult_base) next - have "e1 + \log 2 \real m1\\ - 1 = \log 2 \?a\\ - 1" + have "e1 + \log 2 \real_of_int m1\\ - 1 = \log 2 \?a\\ - 1" using \m1 \ 0\ by (simp add: floor_add2[symmetric] algebra_simps log_mult abs_mult del: floor_add2) also have "\ \ \log 2 \?a + ?b\\" using a_half_less_sum \m1 \ 0\ \?sum \ 0\ - unfolding floor_subtract[symmetric] - by (auto simp add: log_minus_eq_powr powr_minus_divide - intro!: floor_mono) + unfolding floor_diff_of_int[symmetric] + by (auto simp add: log_minus_eq_powr powr_minus_divide intro!: floor_mono) finally have "int p - \log 2 \?a + ?b\\ \ p - (bitlen \m1\) - e1 + 2" by (auto simp: algebra_simps bitlen_def \m1 \ 0\) @@ -1822,11 +1810,11 @@ also have "sgn ?b = sgn m2" using powr_gt_zero[of 2 e2] by (auto simp add: sgn_if zero_less_mult_iff simp del: powr_gt_zero) - also have "\(real (2 * ?m1) + real (sgn m2)) * 2 powr real (?f - - ?e - 1)\ = + also have "\(real_of_int (2 * ?m1) + real_of_int (sgn m2)) * 2 powr real_of_int (?f - - ?e - 1)\ = \Float (?m1 * 2 + sgn m2) (?e - 1) * 2 powr ?f\" by (simp add: powr_add[symmetric] algebra_simps powr_realpow[symmetric]) finally - show "\(?a + ?b) * 2 powr ?f\ = \real (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\" . + show "\(?a + ?b) * 2 powr ?f\ = \real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\" . qed then show ?thesis by transfer (simp add: plus_down_def ac_simps Let_def) @@ -1869,30 +1857,30 @@ subsection \Lemmas needed by Approximate\ lemma Float_num[simp]: - "real (Float 1 0) = 1" - "real (Float 1 1) = 2" - "real (Float 1 2) = 4" - "real (Float 1 (- 1)) = 1/2" - "real (Float 1 (- 2)) = 1/4" - "real (Float 1 (- 3)) = 1/8" - "real (Float (- 1) 0) = -1" - "real (Float (number_of n) 0) = number_of n" + "real_of_float (Float 1 0) = 1" + "real_of_float (Float 1 1) = 2" + "real_of_float (Float 1 2) = 4" + "real_of_float (Float 1 (- 1)) = 1/2" + "real_of_float (Float 1 (- 2)) = 1/4" + "real_of_float (Float 1 (- 3)) = 1/8" + "real_of_float (Float (- 1) 0) = -1" + "real_of_float (Float (number_of n) 0) = number_of n" using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"] two_powr_int_float[of "-3"] using powr_realpow[of 2 2] powr_realpow[of 2 3] using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3] by auto -lemma real_of_Float_int[simp]: "real (Float n 0) = real n" +lemma real_of_Float_int[simp]: "real_of_float (Float n 0) = real n" by simp -lemma float_zero[simp]: "real (Float 0 e) = 0" +lemma float_zero[simp]: "real_of_float (Float 0 e) = 0" by simp lemma abs_div_2_less: "a \ 0 \ a \ -1 \ abs((a::int) div 2) < abs a" by arith -lemma lapprox_rat: "real (lapprox_rat prec x y) \ real x / real y" +lemma lapprox_rat: "real_of_float (lapprox_rat prec x y) \ real_of_int x / real_of_int y" using round_down by (simp add: lapprox_rat_def) lemma mult_div_le: @@ -1914,16 +1902,16 @@ lemma lapprox_rat_nonneg: fixes n x y assumes "0 \ x" and "0 \ y" - shows "0 \ real (lapprox_rat n x y)" + shows "0 \ real_of_float (lapprox_rat n x y)" using assms by (auto simp: lapprox_rat_def simp: round_down_nonneg) -lemma rapprox_rat: "real x / real y \ real (rapprox_rat prec x y)" +lemma rapprox_rat: "real_of_int x / real_of_int y \ real_of_float (rapprox_rat prec x y)" using round_up by (simp add: rapprox_rat_def) lemma rapprox_rat_le1: fixes n x y assumes xy: "0 \ x" "0 < y" "x \ y" - shows "real (rapprox_rat n x y) \ 1" + shows "real_of_float (rapprox_rat n x y) \ 1" proof - have "bitlen \x\ \ bitlen \y\" using xy unfolding bitlen_def by (auto intro!: floor_mono) @@ -1931,10 +1919,10 @@ by transfer (auto intro!: round_up_le1 simp: rat_precision_def) qed -lemma rapprox_rat_nonneg_nonpos: "0 \ x \ y \ 0 \ real (rapprox_rat n x y) \ 0" +lemma rapprox_rat_nonneg_nonpos: "0 \ x \ y \ 0 \ real_of_float (rapprox_rat n x y) \ 0" by transfer (simp add: round_up_le0 divide_nonneg_nonpos) -lemma rapprox_rat_nonpos_nonneg: "x \ 0 \ 0 \ y \ real (rapprox_rat n x y) \ 0" +lemma rapprox_rat_nonpos_nonneg: "x \ 0 \ 0 \ y \ real_of_float (rapprox_rat n x y) \ 0" by transfer (simp add: round_up_le0 divide_nonpos_nonneg) lemma real_divl: "real_divl prec x y \ x / y" @@ -1943,7 +1931,7 @@ lemma real_divr: "x / y \ real_divr prec x y" using round_up by (simp add: real_divr_def) -lemma float_divl: "real (float_divl prec x y) \ real x / real y" +lemma float_divl: "real_of_float (float_divl prec x y) \ x / y" by transfer (rule real_divl) lemma real_divl_lower_bound: @@ -1951,7 +1939,7 @@ by (simp add: real_divl_def round_down_nonneg) lemma float_divl_lower_bound: - "0 \ x \ 0 \ y \ 0 \ real (float_divl prec x y)" + "0 \ x \ 0 \ y \ 0 \ real_of_float (float_divl prec x y)" by transfer (rule real_divl_lower_bound) lemma exponent_1: "exponent 1 = 0" @@ -1968,7 +1956,7 @@ proof show ?rhs if ?lhs proof - - from that have z: "0 = real x" + from that have z: "0 = real_of_float x" using mantissa_exponent by simp show ?thesis by (simp add: zero_float_def z) @@ -1999,17 +1987,17 @@ assumes "0 < x" "x \ 1" "prec \ 1" shows "1 \ real_divl prec 1 x" proof - - have "log 2 x \ real prec + real \log 2 x\" + have "log 2 x \ real_of_int prec + real_of_int \log 2 x\" using \prec \ 1\ by arith from this assms show ?thesis by (simp add: real_divl_def log_divide round_down_ge1) qed lemma float_divl_pos_less1_bound: - "0 < real x \ real x \ 1 \ prec \ 1 \ 1 \ real (float_divl prec 1 x)" + "0 < real_of_float x \ real_of_float x \ 1 \ prec \ 1 \ 1 \ real_of_float (float_divl prec 1 x)" by transfer (rule real_divl_pos_less1_bound) -lemma float_divr: "real x / real y \ real (float_divr prec x y)" +lemma float_divr: "real_of_float x / real_of_float y \ real_of_float (float_divr prec x y)" by transfer (rule real_divr) lemma real_divr_pos_less1_lower_bound: @@ -2032,7 +2020,7 @@ by (simp add: real_divr_def round_up_le0 divide_le_0_iff) lemma float_divr_nonpos_pos_upper_bound: - "real x \ 0 \ 0 \ real y \ real (float_divr prec x y) \ 0" + "real_of_float x \ 0 \ 0 \ real_of_float y \ real_of_float (float_divr prec x y) \ 0" by transfer (rule real_divr_nonpos_pos_upper_bound) lemma real_divr_nonneg_neg_upper_bound: @@ -2040,7 +2028,7 @@ by (simp add: real_divr_def round_up_le0 divide_le_0_iff) lemma float_divr_nonneg_neg_upper_bound: - "0 \ real x \ real y \ 0 \ real (float_divr prec x y) \ 0" + "0 \ real_of_float x \ real_of_float y \ 0 \ real_of_float (float_divr prec x y) \ 0" by transfer (rule real_divr_nonneg_neg_upper_bound) lemma truncate_up_nonneg_mono: @@ -2063,27 +2051,27 @@ have logless: "log 2 x < log 2 y" and flogless: "\log 2 x\ < \log 2 y\" by (metis floor_less_cancel linorder_cases not_le)+ have "truncate_up prec x = - real \x * 2 powr real (int prec - \log 2 x\ - 1)\ * 2 powr - real (int prec - \log 2 x\ - 1)" + real_of_int \x * 2 powr real_of_int (int prec - \log 2 x\ - 1)\ * 2 powr - real_of_int (int prec - \log 2 x\ - 1)" using assms by (simp add: truncate_up_def round_up_def) - also have "\x * 2 powr real (int prec - \log 2 x\ - 1)\ \ (2 ^ prec)" - proof (unfold ceiling_le_eq) - have "x * 2 powr real (int prec - \log 2 x\ - 1) \ x * (2 powr real prec / (2 powr log 2 x))" + also have "\x * 2 powr real_of_int (int prec - \log 2 x\ - 1)\ \ (2 ^ prec)" + proof (unfold ceiling_le_iff) + have "x * 2 powr real_of_int (int prec - \log 2 x\ - 1) \ x * (2 powr real prec / (2 powr log 2 x))" using real_of_int_floor_add_one_ge[of "log 2 x"] assms by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono) - then show "x * 2 powr real (int prec - \log 2 x\ - 1) \ real ((2::int) ^ prec)" + then show "x * 2 powr real_of_int (int prec - \log 2 x\ - 1) \ real_of_int ((2::int) ^ prec)" using \0 < x\ by (simp add: powr_realpow) qed - then have "real \x * 2 powr real (int prec - \log 2 x\ - 1)\ \ 2 powr int prec" + then have "real_of_int \x * 2 powr real_of_int (int prec - \log 2 x\ - 1)\ \ 2 powr int prec" by (auto simp: powr_realpow) also - have "2 powr - real (int prec - \log 2 x\ - 1) \ 2 powr - real (int prec - \log 2 y\)" + have "2 powr - real_of_int (int prec - \log 2 x\ - 1) \ 2 powr - real_of_int (int prec - \log 2 y\)" using logless flogless by (auto intro!: floor_mono) - also have "2 powr real (int prec) \ 2 powr (log 2 y + real (int prec - \log 2 y\))" + also have "2 powr real_of_int (int prec) \ 2 powr (log 2 y + real_of_int (int prec - \log 2 y\))" using assms \0 < x\ by (auto simp: algebra_simps) - finally have "truncate_up prec x \ 2 powr (log 2 y + real (int prec - \log 2 y\)) * 2 powr - real (int prec - \log 2 y\)" + finally have "truncate_up prec x \ 2 powr (log 2 y + real_of_int (int prec - \log 2 y\)) * 2 powr - real_of_int (int prec - \log 2 y\)" by simp - also have "\ = 2 powr (log 2 y + real (int prec - \log 2 y\) - real (int prec - \log 2 y\))" + also have "\ = 2 powr (log 2 y + real_of_int (int prec - \log 2 y\) - real_of_int (int prec - \log 2 y\))" by (subst powr_add[symmetric]) simp also have "\ = y" using \0 < x\ assms @@ -2112,25 +2100,25 @@ assumes "0 < x" "x \ y" shows "truncate_down 0 x \ truncate_down 0 y" proof - - have "x * 2 powr (- real \log 2 x\ - 1) = x * inverse (2 powr ((real \log 2 x\ + 1)))" + have "x * 2 powr (- real_of_int \log 2 x\ - 1) = x * inverse (2 powr ((real_of_int \log 2 x\ + 1)))" by (simp add: powr_divide2[symmetric] powr_add powr_minus inverse_eq_divide) - also have "\ = 2 powr (log 2 x - (real \log 2 x\) - 1)" + also have "\ = 2 powr (log 2 x - (real_of_int \log 2 x\) - 1)" using \0 < x\ by (auto simp: field_simps powr_add powr_divide2[symmetric]) also have "\ < 2 powr 0" using real_of_int_floor_add_one_gt unfolding neg_less_iff_less by (intro powr_less_mono) (auto simp: algebra_simps) - finally have "\x * 2 powr (- real \log 2 x\ - 1)\ < 1" - unfolding less_ceiling_eq real_of_int_minus real_of_one + finally have "\x * 2 powr (- real_of_int \log 2 x\ - 1)\ < 1" + unfolding less_ceiling_iff of_int_minus of_int_1 by simp - moreover have "0 \ \x * 2 powr (- real \log 2 x\ - 1)\" + moreover have "0 \ \x * 2 powr (- real_of_int \log 2 x\ - 1)\" using \x > 0\ by auto - ultimately have "\x * 2 powr (- real \log 2 x\ - 1)\ \ {0 ..< 1}" + ultimately have "\x * 2 powr (- real_of_int \log 2 x\ - 1)\ \ {0 ..< 1}" by simp also have "\ \ {0}" by auto - finally have "\x * 2 powr (- real \log 2 x\ - 1)\ = 0" + finally have "\x * 2 powr (- real_of_int \log 2 x\ - 1)\ = 0" by simp with assms show ?thesis by (auto simp: truncate_down_def round_down_def) @@ -2181,30 +2169,30 @@ by auto have "2 powr (prec - 1) \ y * 2 powr real (prec - 1) / (2 powr log 2 y)" using \0 < y\ by simp - also have "\ \ y * 2 powr real prec / (2 powr (real \log 2 y\ + 1))" + also have "\ \ y * 2 powr real prec / (2 powr (real_of_int \log 2 y\ + 1))" using \0 \ y\ \0 \ x\ assms(2) by (auto intro!: powr_mono divide_left_mono - simp: real_of_nat_diff powr_add + simp: of_nat_diff powr_add powr_divide2[symmetric]) - also have "\ = y * 2 powr real prec / (2 powr real \log 2 y\ * 2)" + also have "\ = y * 2 powr real prec / (2 powr real_of_int \log 2 y\ * 2)" by (auto simp: powr_add) - finally have "(2 ^ (prec - 1)) \ \y * 2 powr real (int prec - \log 2 \y\\ - 1)\" + finally have "(2 ^ (prec - 1)) \ \y * 2 powr real_of_int (int prec - \log 2 \y\\ - 1)\" using \0 \ y\ - by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow) - then have "(2 ^ (prec - 1)) * 2 powr - real (int prec - \log 2 \y\\ - 1) \ truncate_down prec y" + by (auto simp: powr_divide2[symmetric] le_floor_iff powr_realpow) + then have "(2 ^ (prec - 1)) * 2 powr - real_of_int (int prec - \log 2 \y\\ - 1) \ truncate_down prec y" by (auto simp: truncate_down_def round_down_def) moreover { have "x = 2 powr (log 2 \x\)" using \0 < x\ by simp - also have "\ \ (2 ^ (prec )) * 2 powr - real (int prec - \log 2 \x\\ - 1)" + also have "\ \ (2 ^ (prec )) * 2 powr - real_of_int (int prec - \log 2 \x\\ - 1)" using real_of_int_floor_add_one_ge[of "log 2 \x\"] by (auto simp: powr_realpow[symmetric] powr_add[symmetric] algebra_simps) also - have "2 powr - real (int prec - \log 2 \x\\ - 1) \ 2 powr - real (int prec - \log 2 \y\\)" + have "2 powr - real_of_int (int prec - \log 2 \x\\ - 1) \ 2 powr - real_of_int (int prec - \log 2 \y\\)" using logless flogless \x > 0\ \y > 0\ by (auto intro!: floor_mono) - finally have "x \ (2 ^ (prec - 1)) * 2 powr - real (int prec - \log 2 \y\\ - 1)" - by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms real_of_nat_diff) + finally have "x \ (2 ^ (prec - 1)) * 2 powr - real_of_int (int prec - \log 2 \y\\ - 1)" + by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms of_nat_diff) } ultimately show ?thesis by (metis dual_order.trans truncate_down) @@ -2231,12 +2219,12 @@ lemma real_of_float_pprt[simp]: fixes a :: float - shows "real (pprt a) = pprt (real a)" + shows "real_of_float (pprt a) = pprt (real_of_float a)" unfolding pprt_def sup_float_def max_def sup_real_def by auto lemma real_of_float_nprt[simp]: fixes a :: float - shows "real (nprt a) = nprt (real a)" + shows "real_of_float (nprt a) = nprt (real_of_float a)" unfolding nprt_def inf_float_def min_def inf_real_def by auto context @@ -2246,20 +2234,24 @@ qualified lemma compute_int_floor_fl[code]: "int_floor_fl (Float m e) = (if 0 \ e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))" - by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints) + apply transfer + apply (simp add: powr_int floor_divide_of_int_eq) + by (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power floor_of_int of_int_mult) -lift_definition floor_fl :: "float \ float" is "\x. real (floor x)" by simp +lift_definition floor_fl :: "float \ float" is "\x. real_of_int (floor x)" by simp qualified lemma compute_floor_fl[code]: "floor_fl (Float m e) = (if 0 \ e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)" - by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints) + apply transfer + apply (simp add: powr_int floor_divide_of_int_eq) + by (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power of_int_mult) end -lemma floor_fl: "real (floor_fl x) \ real x" +lemma floor_fl: "real_of_float (floor_fl x) \ real_of_float x" by transfer simp -lemma int_floor_fl: "real (int_floor_fl x) \ real x" +lemma int_floor_fl: "real_of_int (int_floor_fl x) \ real_of_float x" by transfer simp lemma floor_pos_exp: "exponent (floor_fl x) \ 0" @@ -2269,9 +2261,9 @@ by (simp add: floor_fl_def) next case False - have eq: "floor_fl x = Float \real x\ 0" + have eq: "floor_fl x = Float \real_of_float x\ 0" by transfer simp - obtain i where "\real x\ = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i" + obtain i where "\real_of_float x\ = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i" by (rule denormalize_shift[OF eq[THEN eq_reflection] False]) then show ?thesis by simp diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Library/Formal_Power_Series.thy --- a/src/HOL/Library/Formal_Power_Series.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Library/Formal_Power_Series.thy Tue Nov 10 14:18:41 2015 +0000 @@ -3143,8 +3143,8 @@ "setsum (\k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n" using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n] apply (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric] - of_nat_setsum[symmetric] of_nat_add[symmetric]) - apply simp + of_nat_setsum[symmetric] of_nat_add[symmetric]) + apply blast done lemma binomial_Vandermonde_same: "setsum (\k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Library/positivstellensatz.ML --- a/src/HOL/Library/positivstellensatz.ML Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Library/positivstellensatz.ML Tue Nov 10 14:18:41 2015 +0000 @@ -695,8 +695,8 @@ fun is_alien ct = case Thm.term_of ct of - Const(@{const_name "real"}, _)$ n => - if can HOLogic.dest_number n then false else true + Const(@{const_name "of_nat"}, _)$ n => not (can HOLogic.dest_number n) + | Const(@{const_name "of_int"}, _)$ n => not (can HOLogic.dest_number n) | _ => false in fun real_linear_prover translator (eq,le,lt) = @@ -711,7 +711,7 @@ (eq_pols @ le_pols @ lt_pols) []) val le_pols' = le_pols @ map (fn v => FuncUtil.Ctermfunc.onefunc (v,Rat.one)) aliens val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols) - val le' = le @ map (fn a => Thm.instantiate' [] [SOME (Thm.dest_arg a)] @{thm real_of_nat_ge_zero}) aliens + val le' = le @ map (fn a => Thm.instantiate' [] [SOME (Thm.dest_arg a)] @{thm of_nat_0_le_iff}) aliens in ((translator (eq,le',lt) proof), Trivial) end end; diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Limits.thy --- a/src/HOL/Limits.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Limits.thy Tue Nov 10 14:18:41 2015 +0000 @@ -175,7 +175,7 @@ ultimately have "\m. norm (X m) \ real (Suc n)" by (blast intro: order_trans) then show "\N. \n. norm (X n) \ real (Suc N)" .. -qed (force simp add: real_of_nat_Suc) +qed (force simp add: of_nat_Suc) text\alternative definition for Bseq\ lemma Bseq_iff: "Bseq X = (\N. \n. norm (X n) \ real(Suc N))" @@ -188,7 +188,7 @@ apply (subst lemma_NBseq_def, auto) apply (rule_tac x = "Suc N" in exI) apply (rule_tac [2] x = N in exI) -apply (auto simp add: real_of_nat_Suc) +apply (auto simp add: of_nat_Suc) prefer 2 apply (blast intro: order_less_imp_le) apply (drule_tac x = n in spec, simp) done @@ -1433,7 +1433,7 @@ proof (induct n) case (Suc n) have "real (Suc n) * x + 1 \ (x + 1) * (real n * x + 1)" - by (simp add: field_simps real_of_nat_Suc x) + by (simp add: field_simps of_nat_Suc x) also have "\ \ (x + 1)^Suc n" using Suc x by (simp add: mult_left_mono) finally show ?case . diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/MacLaurin.thy --- a/src/HOL/MacLaurin.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/MacLaurin.thy Tue Nov 10 14:18:41 2015 +0000 @@ -24,17 +24,17 @@ lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))" by arith -lemma fact_diff_Suc [rule_format]: - "n < Suc m ==> fact (Suc m - n) = (Suc m - n) * fact (m - n)" +lemma fact_diff_Suc: + "n < Suc m \ fact (Suc m - n) = (Suc m - n) * fact (m - n)" by (subst fact_reduce, auto) lemma Maclaurin_lemma2: fixes B assumes DERIV : "\m t. m < n \ 0\t \ t\h \ DERIV (diff m) t :> diff (Suc m) t" and INIT : "n = Suc k" - defines "difg \ - (\m t::real. diff m t - - ((\p + (\m t::real. diff m t - + ((\p (\m t. diff m t - ?difg m t)") shows "\m t. m < n & 0 \ t & t \ h --> DERIV (difg m) t :> difg (Suc m) t" proof (rule allI impI)+ @@ -42,10 +42,9 @@ assume INIT2: "m < n & 0 \ t & t \ h" have "DERIV (difg m) t :> diff (Suc m) t - ((\xx. (fact x) + real x * (fact x) \ 0" - by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2 less_numeral_extra(3) mult_less_0_iff not_real_of_nat_less_zero) + by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2 less_numeral_extra(3) mult_less_0_iff of_nat_less_0_iff) have "\x. (Suc x) * t ^ x * diff (Suc m + x) 0 / (fact (Suc x)) = diff (Suc m + x) 0 * t^x / (fact x)" by (rule nonzero_divide_eq_eq[THEN iffD2]) auto @@ -64,7 +63,7 @@ using \0 < n - m\ by (simp add: divide_simps fact_reduce) ultimately show "DERIV (difg m) t :> difg (Suc m) t" - unfolding difg_def by simp + unfolding difg_def by (simp add: mult.commute) qed lemma Maclaurin: @@ -420,7 +419,7 @@ lemma sin_expansion_lemma: "sin (x + real (Suc m) * pi / 2) = cos (x + real (m) * pi / 2)" -by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib distrib_right, auto) +by (simp only: cos_add sin_add of_nat_Suc add_divide_distrib distrib_right, auto) lemma Maclaurin_sin_expansion2: "\t. abs t \ abs x & @@ -431,7 +430,7 @@ and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl) apply safe apply (simp) - apply (simp add: sin_expansion_lemma del: real_of_nat_Suc) + apply (simp add: sin_expansion_lemma del: of_nat_Suc) apply (force intro!: derivative_eq_intros) apply (subst (asm) setsum.neutral, auto)[1] apply (rule ccontr, simp) @@ -439,7 +438,7 @@ apply (erule ssubst) apply (rule_tac x = t in exI, simp) apply (rule setsum.cong[OF refl]) -apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: real_of_nat_Suc) +apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc) done lemma Maclaurin_sin_expansion: @@ -459,12 +458,12 @@ apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl) apply safe apply simp - apply (simp (no_asm) add: sin_expansion_lemma del: real_of_nat_Suc) + apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc) apply (force intro!: derivative_eq_intros) apply (erule ssubst) apply (rule_tac x = t in exI, simp) apply (rule setsum.cong[OF refl]) - apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: real_of_nat_Suc) + apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc) done lemma Maclaurin_sin_expansion4: @@ -476,12 +475,12 @@ apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl) apply safe apply simp - apply (simp (no_asm) add: sin_expansion_lemma del: real_of_nat_Suc) + apply (simp (no_asm) add: sin_expansion_lemma del: of_nat_Suc) apply (force intro!: derivative_eq_intros) apply (erule ssubst) apply (rule_tac x = t in exI, simp) apply (rule setsum.cong[OF refl]) - apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: real_of_nat_Suc) + apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc) done @@ -493,7 +492,7 @@ lemma cos_expansion_lemma: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)" -by (simp only: cos_add sin_add real_of_nat_Suc distrib_right add_divide_distrib, auto) +by (simp only: cos_add sin_add of_nat_Suc distrib_right add_divide_distrib, auto) lemma Maclaurin_cos_expansion: "\t::real. abs t \ abs x & @@ -503,7 +502,7 @@ apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl) apply safe apply (simp (no_asm)) - apply (simp (no_asm) add: cos_expansion_lemma del: real_of_nat_Suc) + apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc) apply (case_tac "n", simp) apply (simp del: setsum_lessThan_Suc) apply (rule ccontr, simp) @@ -523,7 +522,7 @@ apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl) apply safe apply simp - apply (simp (no_asm) add: cos_expansion_lemma del: real_of_nat_Suc) + apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc) apply (erule ssubst) apply (rule_tac x = t in exI, simp) apply (rule setsum.cong[OF refl]) @@ -539,7 +538,7 @@ apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl) apply safe apply simp - apply (simp (no_asm) add: cos_expansion_lemma del: real_of_nat_Suc) + apply (simp (no_asm) add: cos_expansion_lemma del: of_nat_Suc) apply (erule ssubst) apply (rule_tac x = t in exI, simp) apply (rule setsum.cong[OF refl]) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Matrix_LP/ComputeFloat.thy --- a/src/HOL/Matrix_LP/ComputeFloat.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Matrix_LP/ComputeFloat.thy Tue Nov 10 14:18:41 2015 +0000 @@ -10,22 +10,24 @@ ML_file "~~/src/Tools/float.ML" +(*FIXME surely floor should be used? This file is full of redundant material.*) + definition int_of_real :: "real \ int" - where "int_of_real x = (SOME y. real y = x)" + where "int_of_real x = (SOME y. real_of_int y = x)" definition real_is_int :: "real \ bool" - where "real_is_int x = (EX (u::int). x = real u)" + where "real_is_int x = (EX (u::int). x = real_of_int u)" -lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))" +lemma real_is_int_def2: "real_is_int x = (x = real_of_int (int_of_real x))" by (auto simp add: real_is_int_def int_of_real_def) -lemma real_is_int_real[simp]: "real_is_int (real (x::int))" +lemma real_is_int_real[simp]: "real_is_int (real_of_int (x::int))" by (auto simp add: real_is_int_def int_of_real_def) -lemma int_of_real_real[simp]: "int_of_real (real x) = x" +lemma int_of_real_real[simp]: "int_of_real (real_of_int x) = x" by (simp add: int_of_real_def) -lemma real_int_of_real[simp]: "real_is_int x \ real (int_of_real x) = x" +lemma real_int_of_real[simp]: "real_is_int x \ real_of_int (int_of_real x) = x" by (auto simp add: int_of_real_def real_is_int_def) lemma real_is_int_add_int_of_real: "real_is_int a \ real_is_int b \ (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)" @@ -44,15 +46,15 @@ apply (simp add: int_of_real_sub real_int_of_real) done -lemma real_is_int_rep: "real_is_int x \ ?! (a::int). real a = x" +lemma real_is_int_rep: "real_is_int x \ ?! (a::int). real_of_int a = x" by (auto simp add: real_is_int_def) lemma int_of_real_mult: assumes "real_is_int a" "real_is_int b" shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)" using assms - by (auto simp add: real_is_int_def real_of_int_mult[symmetric] - simp del: real_of_int_mult) + by (auto simp add: real_is_int_def of_int_mult[symmetric] + simp del: of_int_mult) lemma real_is_int_mult[simp]: "real_is_int a \ real_is_int b \ real_is_int (a*b)" apply (subst real_is_int_def2) @@ -64,14 +66,14 @@ lemma real_is_int_1[simp]: "real_is_int (1::real)" proof - - have "real_is_int (1::real) = real_is_int(real (1::int))" by auto + have "real_is_int (1::real) = real_is_int(real_of_int (1::int))" by auto also have "\ = True" by (simp only: real_is_int_real) ultimately show ?thesis by auto qed lemma real_is_int_n1: "real_is_int (-1::real)" proof - - have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto + have "real_is_int (-1::real) = real_is_int(real_of_int (-1::int))" by auto also have "\ = True" by (simp only: real_is_int_real) ultimately show ?thesis by auto qed @@ -87,19 +89,16 @@ lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)" proof - - have 1: "(1::real) = real (1::int)" by auto + have 1: "(1::real) = real_of_int (1::int)" by auto show ?thesis by (simp only: 1 int_of_real_real) qed lemma int_of_real_numeral[simp]: "int_of_real (numeral b) = numeral b" - unfolding int_of_real_def - by (intro some_equality) - (auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject) + unfolding int_of_real_def by simp lemma int_of_real_neg_numeral[simp]: "int_of_real (- numeral b) = - numeral b" unfolding int_of_real_def - by (intro some_equality) - (auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject) + by (metis int_of_real_def int_of_real_real of_int_minus of_int_of_nat_eq of_nat_numeral) lemma int_div_zdiv: "int (a div b) = (int a) div (int b)" by (rule zdiv_int) @@ -160,7 +159,7 @@ zpower_numeral_odd zpower_Pls definition float :: "(int \ int) \ real" where - "float = (\(a, b). real a * 2 powr real b)" + "float = (\(a, b). real_of_int a * 2 powr real_of_int b)" lemma float_add_l0: "float (0, e) + x = x" by (simp add: float_def) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Matrix_LP/ComputeNumeral.thy --- a/src/HOL/Matrix_LP/ComputeNumeral.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Matrix_LP/ComputeNumeral.thy Tue Nov 10 14:18:41 2015 +0000 @@ -14,7 +14,7 @@ (* addition for bit strings *) lemmas bitadd = add_num_simps -(* multiplication for bit strings *) +(* multiplication for bit strings *) lemmas bitmul = mult_num_simps lemmas bitarith = arith_simps @@ -38,9 +38,9 @@ arith_simps rel_simps number_norm lemmas compute_num_conversions = - real_of_nat_numeral real_of_nat_zero + of_nat_numeral of_nat_0 nat_numeral nat_0 nat_neg_numeral - real_numeral real_of_int_zero + of_int_numeral of_int_neg_numeral of_int_0 lemmas zpowerarith = zpower_numeral_even zpower_numeral_odd zpower_Pls int_pow_1 @@ -74,7 +74,7 @@ lemmas compute_even = even_0_int even_One_int even_Bit0_int even_Bit1_int -lemmas compute_numeral = compute_if compute_let compute_pair compute_bool +lemmas compute_numeral = compute_if compute_let compute_pair compute_bool compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even end diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy Tue Nov 10 14:18:41 2015 +0000 @@ -48,7 +48,7 @@ lemma swap_apply1: "Fun.swap x y f x = f y" by (simp add: Fun.swap_def) - + lemma swap_apply2: "Fun.swap x y f y = f x" by (simp add: Fun.swap_def) @@ -181,7 +181,7 @@ shows "odd (card {s\simplices. (rl ` s = {..Suc n})})" proof (rule kuhn_counting_lemma) have finite_s[simp]: "\s. s \ simplices \ finite s" - by (metis add_is_0 zero_neq_numeral card_infinite assms(3)) + by (metis add_is_0 zero_neq_numeral card_infinite assms(3)) let ?F = "{f. \s\simplices. face f s}" have F_eq: "?F = (\s\simplices. \a\s. {s - {a}})" @@ -293,7 +293,7 @@ proof - { fix x y :: nat assume "x \ y" "x \ n" "y \ n" with upd have "upd ` {..< x} \ upd ` {..< y}" - by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def) + by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def) then have "enum x \ enum y" by (auto simp add: enum_def fun_eq_iff) } then show ?thesis @@ -446,7 +446,7 @@ using step \j + d \ n\ eq by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" using \j + d \ n\ - by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) + by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) (auto intro!: s.enum_in t.enum_in) then show ?case by simp qed @@ -466,7 +466,7 @@ moreover have e0: "a = s.enum 0" "b = t.enum 0" using a b by (simp_all add: s.enum_0_bot t.enum_0_bot) moreover - { fix j assume "0 < j" "j \ n" + { fix j assume "0 < j" "j \ n" moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}" unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum j = t.enum j" @@ -505,7 +505,7 @@ moreover have en: "a = s.enum n" "b = t.enum n" using a b by (simp_all add: s.enum_n_top t.enum_n_top) moreover - { fix j assume "j < n" + { fix j assume "j < n" moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}" unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc) ultimately have "s.enum j = t.enum j" @@ -665,7 +665,7 @@ proof cases case (ksimplex b_s u_s) - { fix t b assume "ksimplex p n t" + { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t @@ -688,7 +688,7 @@ proof cases case (ksimplex b_s u_s) - { fix t b assume "ksimplex p n t" + { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t @@ -858,7 +858,7 @@ have ks_f': "ksimplex p n (b.enum ` {.. n})" unfolding f' by rule unfold_locales - have "0 < n" + have "0 < n" using \n \ 0\ by auto { from \a = enum i\ \n \ 0\ \i = n\ lb upd_space[of n'] @@ -1040,7 +1040,7 @@ using i by (simp add: k i'_def) also have "\ = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" using \Suc l < k\ \k \ n\ by (simp add: t.enum_Suc l t.upd_inj) - finally have "(u l = upd i' \ u (Suc l) = upd (Suc i')) \ + finally have "(u l = upd i' \ u (Suc l) = upd (Suc i')) \ (u l = upd (Suc i') \ u (Suc l) = upd i')" using \Suc i' < n\ by (auto simp: enum_Suc fun_eq_iff split: split_if_asm) @@ -1187,7 +1187,7 @@ then have "reduced (Suc n) (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover - have "n \ (reduced (Suc n) \ lab) ` f" + have "n \ (reduced (Suc n) \ lab) ` f" using eq by auto ultimately show False by force @@ -1657,7 +1657,7 @@ have q2: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0)" unfolding o_def using cube \p > 0\ by (intro allI impI label(2)) (auto simp add: b'') - have q3: "\x. (\i p) \ (\i + have q3: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" using cube \p > 0\ unfolding o_def by (intro allI impI label(3)) (auto simp add: b'') obtain q where q: @@ -1694,7 +1694,7 @@ done also have "\ = d" using DIM_positive[where 'a='a] - by (auto simp: real_eq_of_nat n_def) + by (auto simp: n_def) finally show False using d_fz_z by auto qed @@ -1742,11 +1742,11 @@ have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" apply (rule setsum_mono) using rs(1)[OF b'_im] - apply (auto simp add:* field_simps simp del: real_of_nat_Suc) + apply (auto simp add:* field_simps simp del: of_nat_Suc) done also have "\ < e * real p" using p \e > 0\ \p > 0\ - by (auto simp add: field_simps n_def real_of_nat_def) + by (auto simp add: field_simps n_def) finally have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) < e * real p" . } moreover @@ -1754,11 +1754,11 @@ have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" apply (rule setsum_mono) using rs(2)[OF b'_im] - apply (auto simp add:* field_simps simp del: real_of_nat_Suc) + apply (auto simp add:* field_simps simp del: of_nat_Suc) done also have "\ < e * real p" using p \e > 0\ \p > 0\ - by (auto simp add: field_simps n_def real_of_nat_def) + by (auto simp add: field_simps n_def) finally have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) < e * real p" . } ultimately diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Tue Nov 10 14:18:41 2015 +0000 @@ -2986,17 +2986,17 @@ } note ptgh_ee = this have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \ ball (p t) (ee (p t))" using ptgh_ee [of "n/N"] Suc.prems - by (auto simp: field_simps real_of_nat_def dist_norm dest: segment_furthest_le [where y="p t"]) + by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ s" using \N>0\ Suc.prems - apply (simp add: real_of_nat_def path_image_join field_simps closed_segment_commute) + apply (simp add: path_image_join field_simps closed_segment_commute) apply (erule order_trans) apply (simp add: ee pi t) done have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) \ ball (p t) (ee (p t))" using ptgh_ee [of "(1+n)/N"] Suc.prems - by (auto simp: field_simps real_of_nat_def dist_norm dest: segment_furthest_le [where y="p t"]) + by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ s" using \N>0\ Suc.prems ee pi t by (auto simp: Path_Connected.path_image_join field_simps) @@ -3030,7 +3030,7 @@ path_integral (linepath (h (n/N)) (g (n/N))) f = 0" using path_integral_unique [OF pi0] Suc.prems by (simp add: g h fpa valid_path_subpath path_integrable_subpath - fpa1 fpa2 fpa3 algebra_simps del: real_of_nat_Suc) + fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. \hn - gn = ghn - gh0; gd + ghn' + he + hgn = (0::complex); @@ -4368,10 +4368,10 @@ subsection{* Cauchy's integral formula, again for a convex enclosing set.*} -lemma Cauchy_integral_formula_weak: - assumes s: "convex s" and "finite k" and conf: "continuous_on s f" - and fcd: "(\x. x \ interior s - k \ f complex_differentiable at x)" - and z: "z \ interior s - k" and vpg: "valid_path \" +lemma Cauchy_integral_formula_weak: + assumes s: "convex s" and "finite k" and conf: "continuous_on s f" + and fcd: "(\x. x \ interior s - k \ f complex_differentiable at x)" + and z: "z \ interior s - k" and vpg: "valid_path \" and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_path_integral (2*pi * ii * winding_number \ z * f z)) \" proof - @@ -4383,7 +4383,7 @@ case True then show ?thesis apply (simp add: continuous_within) apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) - using has_field_derivative_at_within DERIV_within_iff f' + using has_field_derivative_at_within DERIV_within_iff f' apply (fastforce simp add:)+ done next @@ -4395,7 +4395,7 @@ apply (rule continuous_transform_within [OF dxz that, of "\w::complex. (f w - f z)/(w - z)"]) apply (force simp: dist_commute) apply (rule cf continuous_intros)+ - using False by auto + using False by auto qed have fink': "finite (insert z k)" using \finite k\ by blast have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_path_integral 0) \" @@ -4412,9 +4412,9 @@ using znotin has_path_integral_add [OF has_path_integral_lmul [OF has_path_integral_winding_number [OF vpg znotin], of "f z"] *] apply (auto simp: mult_ac divide_simps) done -qed - -theorem Cauchy_integral_formula_convex_simple: +qed + +theorem Cauchy_integral_formula_convex_simple: "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; pathfinish \ = pathstart \\ \ ((\w. f w / (w - z)) has_path_integral (2*pi * ii * winding_number \ z * f z)) \" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy Tue Nov 10 14:18:41 2015 +0000 @@ -19,7 +19,7 @@ shows "((op * c) has_derivative (op * c)) F" by (rule has_derivative_mult_right [OF has_derivative_id]) -lemma has_derivative_of_real[derivative_intros, simp]: +lemma has_derivative_of_real[derivative_intros, simp]: "(f has_derivative f') F \ ((\x. of_real (f x)) has_derivative (\x. of_real (f' x))) F" using bounded_linear.has_derivative[OF bounded_linear_of_real] . @@ -45,39 +45,39 @@ using bounded_linear.linear[OF bounded_linear_cnj] . lemma tendsto_mult_left: - fixes c::"'a::real_normed_algebra" + fixes c::"'a::real_normed_algebra" shows "(f ---> l) F \ ((\x. c * (f x)) ---> c * l) F" by (rule tendsto_mult [OF tendsto_const]) lemma tendsto_mult_right: - fixes c::"'a::real_normed_algebra" + fixes c::"'a::real_normed_algebra" shows "(f ---> l) F \ ((\x. (f x) * c) ---> l * c) F" by (rule tendsto_mult [OF _ tendsto_const]) lemma tendsto_Re_upper: - assumes "~ (trivial_limit F)" - "(f ---> l) F" + assumes "~ (trivial_limit F)" + "(f ---> l) F" "eventually (\x. Re(f x) \ b) F" shows "Re(l) \ b" by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re) lemma tendsto_Re_lower: - assumes "~ (trivial_limit F)" - "(f ---> l) F" + assumes "~ (trivial_limit F)" + "(f ---> l) F" "eventually (\x. b \ Re(f x)) F" shows "b \ Re(l)" by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re) lemma tendsto_Im_upper: - assumes "~ (trivial_limit F)" - "(f ---> l) F" + assumes "~ (trivial_limit F)" + "(f ---> l) F" "eventually (\x. Im(f x) \ b) F" shows "Im(l) \ b" by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im) lemma tendsto_Im_lower: - assumes "~ (trivial_limit F)" - "(f ---> l) F" + assumes "~ (trivial_limit F)" + "(f ---> l) F" "eventually (\x. b \ Im(f x)) F" shows "b \ Im(l)" by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im) @@ -89,29 +89,29 @@ by auto lemma continuous_mult_left: - fixes c::"'a::real_normed_algebra" + fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. c * f x)" by (rule continuous_mult [OF continuous_const]) lemma continuous_mult_right: - fixes c::"'a::real_normed_algebra" + fixes c::"'a::real_normed_algebra" shows "continuous F f \ continuous F (\x. f x * c)" by (rule continuous_mult [OF _ continuous_const]) lemma continuous_on_mult_left: - fixes c::"'a::real_normed_algebra" + fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. c * f x)" by (rule continuous_on_mult [OF continuous_on_const]) lemma continuous_on_mult_right: - fixes c::"'a::real_normed_algebra" + fixes c::"'a::real_normed_algebra" shows "continuous_on s f \ continuous_on s (\x. f x * c)" by (rule continuous_on_mult [OF _ continuous_on_const]) lemma uniformly_continuous_on_cmul_right [continuous_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "uniformly_continuous_on s f \ uniformly_continuous_on s (\x. f x * c)" - using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . + using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . lemma uniformly_continuous_on_cmul_left[continuous_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_algebra" @@ -141,7 +141,7 @@ lemma DERIV_zero_constant: fixes f :: "'a::{real_normed_field, real_inner} \ 'a" shows "\convex s; - \x. x\s \ (f has_field_derivative 0) (at x within s)\ + \x. x\s \ (f has_field_derivative 0) (at x within s)\ \ \c. \x \ s. f(x) = c" by (auto simp: has_field_derivative_def lambda_zero intro: has_derivative_zero_constant) @@ -162,7 +162,7 @@ and d0: "\x. x\s \ DERIV f x :> 0" and "a \ s" and "x \ s" - shows "f x = f a" + shows "f x = f a" by (rule has_derivative_zero_unique_connected [OF assms(2,1) _ assms(5,4)]) (metis has_field_derivative_def lambda_zero d0) @@ -192,7 +192,7 @@ (*generalising DERIV_isconst_all, which requires type real (using the ordering)*) lemma DERIV_zero_UNIV_unique: fixes f :: "'a::{real_normed_field, real_inner} \ 'a" - shows "(\x. DERIV f x :> 0) \ f x = f a" + shows "(\x. DERIV f x :> 0) \ f x = f a" by (metis DERIV_zero_unique UNIV_I assms convex_UNIV) subsection \Some limit theorems about real part of real series etc.\ @@ -263,7 +263,7 @@ shows "(f ---> l) sequentially \ (\N. \n\N. f n \ \) \ l \ \" by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) -lemma real_series: +lemma real_series: fixes l::complex shows "f sums l \ (\n. f n \ \) \ l \ \" unfolding sums_def @@ -276,7 +276,7 @@ subsection\Holomorphic functions\ definition complex_differentiable :: "[complex \ complex, complex filter] \ bool" - (infixr "(complex'_differentiable)" 50) + (infixr "(complex'_differentiable)" 50) where "f complex_differentiable F \ \f'. (f has_field_derivative f') F" lemma complex_differentiable_imp_continuous_at: @@ -304,7 +304,7 @@ lemma complex_differentiable_const [derivative_intros]: "(\z. c) complex_differentiable F" unfolding complex_differentiable_def has_field_derivative_def by (rule exI [where x=0]) - (metis has_derivative_const lambda_zero) + (metis has_derivative_const lambda_zero) lemma complex_differentiable_ident [derivative_intros]: "(\z. z) complex_differentiable F" unfolding complex_differentiable_def has_field_derivative_def @@ -343,14 +343,14 @@ by (metis DERIV_inverse_fun) lemma complex_differentiable_mult [derivative_intros]: - assumes "f complex_differentiable (at a within s)" + assumes "f complex_differentiable (at a within s)" "g complex_differentiable (at a within s)" shows "(\z. f z * g z) complex_differentiable (at a within s)" using assms unfolding complex_differentiable_def by (metis DERIV_mult [of f _ a s g]) - + lemma complex_differentiable_divide [derivative_intros]: - assumes "f complex_differentiable (at a within s)" + assumes "f complex_differentiable (at a within s)" "g complex_differentiable (at a within s)" "g a \ 0" shows "(\z. f z / g z) complex_differentiable (at a within s)" @@ -358,7 +358,7 @@ by (metis DERIV_divide [of f _ a s g]) lemma complex_differentiable_power [derivative_intros]: - assumes "f complex_differentiable (at a within s)" + assumes "f complex_differentiable (at a within s)" shows "(\z. f z ^ n) complex_differentiable (at a within s)" using assms unfolding complex_differentiable_def by (metis DERIV_power) @@ -373,7 +373,7 @@ by (blast intro: has_derivative_transform_within) lemma complex_differentiable_compose_within: - assumes "f complex_differentiable (at a within s)" + assumes "f complex_differentiable (at a within s)" "g complex_differentiable (at (f a) within f`s)" shows "(g o f) complex_differentiable (at a within s)" using assms unfolding complex_differentiable_def @@ -385,7 +385,7 @@ by (metis complex_differentiable_at_within complex_differentiable_compose_within) lemma complex_differentiable_within_open: - "\a \ s; open s\ \ f complex_differentiable at a within s \ + "\a \ s; open s\ \ f complex_differentiable at a within s \ f complex_differentiable at a" unfolding complex_differentiable_def by (metis at_within_open) @@ -409,7 +409,7 @@ definition holomorphic_on :: "[complex \ complex, complex set] \ bool" (infixl "(holomorphic'_on)" 50) where "f holomorphic_on s \ \x\s. f complex_differentiable (at x within s)" - + named_theorems holomorphic_intros "structural introduction rules for holomorphic_on" lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}" @@ -419,9 +419,9 @@ "open s \ f holomorphic_on s \ (\x \ s. \f'. DERIV f x :> f')" by (auto simp: holomorphic_on_def complex_differentiable_def has_field_derivative_def at_within_open [of _ s]) -lemma holomorphic_on_imp_continuous_on: +lemma holomorphic_on_imp_continuous_on: "f holomorphic_on s \ continuous_on s f" - by (metis complex_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) + by (metis complex_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) lemma holomorphic_on_subset: "f holomorphic_on s \ t \ s \ f holomorphic_on t" @@ -505,19 +505,19 @@ by (metis DERIV_imp_deriv DERIV_const) lemma complex_derivative_add: - "\f complex_differentiable at z; g complex_differentiable at z\ + "\f complex_differentiable at z; g complex_differentiable at z\ \ deriv (\w. f w + g w) z = deriv f z + deriv g z" unfolding DERIV_deriv_iff_complex_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_intros) lemma complex_derivative_diff: - "\f complex_differentiable at z; g complex_differentiable at z\ + "\f complex_differentiable at z; g complex_differentiable at z\ \ deriv (\w. f w - g w) z = deriv f z - deriv g z" unfolding DERIV_deriv_iff_complex_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_intros) lemma complex_derivative_mult: - "\f complex_differentiable at z; g complex_differentiable at z\ + "\f complex_differentiable at z; g complex_differentiable at z\ \ deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" unfolding DERIV_deriv_iff_complex_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_eq_intros) @@ -533,7 +533,7 @@ by (auto intro!: DERIV_imp_deriv derivative_eq_intros) lemma complex_derivative_transform_within_open: - "\f holomorphic_on s; g holomorphic_on s; open s; z \ s; \w. w \ s \ f w = g w\ + "\f holomorphic_on s; g holomorphic_on s; open s; z \ s; \w. w \ s \ f w = g w\ \ deriv f z = deriv g z" unfolding holomorphic_on_def by (rule DERIV_imp_deriv) @@ -549,7 +549,7 @@ subsection\Analyticity on a set\ -definition analytic_on (infixl "(analytic'_on)" 50) +definition analytic_on (infixl "(analytic'_on)" 50) where "f analytic_on s \ \x \ s. \e. 0 < e \ f holomorphic_on (ball x e)" @@ -578,7 +578,7 @@ lemma analytic_on_UN: "f analytic_on (\i\I. s i) \ (\i\I. f analytic_on (s i))" by (auto simp: analytic_on_def) - + lemma analytic_on_holomorphic: "f analytic_on s \ (\t. open t \ s \ t \ f holomorphic_on t)" (is "?lhs = ?rhs") @@ -593,7 +593,7 @@ by (metis analytic_on_def) next fix t - assume "open t" "s \ t" "f analytic_on t" + assume "open t" "s \ t" "f analytic_on t" then show "f analytic_on s" by (metis analytic_on_subset) qed @@ -625,17 +625,17 @@ then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g - by (metis analytic_on_def g image_eqI x) + by (metis analytic_on_def g image_eqI x) have "isCont f x" by (metis analytic_on_imp_differentiable_at complex_differentiable_imp_continuous_at f x) with e' obtain d where d: "0 < d" and fd: "f ` ball x d \ ball (f x) e'" by (auto simp: continuous_at_ball) - have "g \ f holomorphic_on ball x (min d e)" + have "g \ f holomorphic_on ball x (min d e)" apply (rule holomorphic_on_compose) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) then show "\e>0. g \ f holomorphic_on ball x e" - by (metis d e min_less_iff_conj) + by (metis d e min_less_iff_conj) qed lemma analytic_on_compose_gen: @@ -658,9 +658,9 @@ then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g - by (metis analytic_on_def g z) - have "(\z. f z + g z) holomorphic_on ball z (min e e')" - apply (rule holomorphic_on_add) + by (metis analytic_on_def g z) + have "(\z. f z + g z) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_add) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z + g z) holomorphic_on ball z e" @@ -678,9 +678,9 @@ then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g - by (metis analytic_on_def g z) - have "(\z. f z - g z) holomorphic_on ball z (min e e')" - apply (rule holomorphic_on_diff) + by (metis analytic_on_def g z) + have "(\z. f z - g z) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_diff) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z - g z) holomorphic_on ball z e" @@ -698,9 +698,9 @@ then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g - by (metis analytic_on_def g z) - have "(\z. f z * g z) holomorphic_on ball z (min e e')" - apply (rule holomorphic_on_mult) + by (metis analytic_on_def g z) + have "(\z. f z * g z) holomorphic_on ball z (min e e')" + apply (rule holomorphic_on_mult) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z * g z) holomorphic_on ball z e" @@ -719,12 +719,12 @@ by (metis analytic_on_def) have "continuous_on (ball z e) f" by (metis fh holomorphic_on_imp_continuous_on) - then obtain e' where e': "0 < e'" and nz': "\y. dist z y < e' \ f y \ 0" - by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz) - have "(\z. inverse (f z)) holomorphic_on ball z (min e e')" + then obtain e' where e': "0 < e'" and nz': "\y. dist z y < e' \ f y \ 0" + by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz) + have "(\z. inverse (f z)) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_inverse) apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) - by (metis nz' mem_ball min_less_iff_conj) + by (metis nz' mem_ball min_less_iff_conj) then show "\e>0. (\z. inverse (f z)) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed @@ -764,9 +764,9 @@ "f analytic_on {z} \ g analytic_on {z} \ (\s. open s \ z \ s \ f holomorphic_on s \ g holomorphic_on s)" (is "?lhs = ?rhs") -proof +proof assume ?lhs - then obtain s t + then obtain s t where st: "open s" "z \ s" "f holomorphic_on s" "open t" "z \ t" "g holomorphic_on t" by (auto simp: analytic_at) @@ -776,14 +776,14 @@ apply (auto simp: Diff_subset holomorphic_on_subset) done next - assume ?rhs + assume ?rhs then show ?lhs by (force simp add: analytic_at) qed subsection\Combining theorems for derivative with ``analytic at'' hypotheses\ -lemma +lemma assumes "f analytic_on {z}" "g analytic_on {z}" shows complex_derivative_add_at: "deriv (\w. f w + g w) z = deriv f z + deriv g z" and complex_derivative_diff_at: "deriv (\w. f w - g w) z = deriv f z - deriv g z" @@ -826,14 +826,14 @@ and df: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x within s)" and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ s \ norm (f' n x - g' x) \ e" and "\x l. x \ s \ ((\n. f n x) ---> l) sequentially" - shows "\g. \x \ s. ((\n. f n x) ---> g x) sequentially \ + shows "\g. \x \ s. ((\n. f n x) ---> g x) sequentially \ (g has_field_derivative (g' x)) (at x within s)" proof - from assms obtain x l where x: "x \ s" and tf: "((\n. f n x) ---> l) sequentially" by blast { fix e::real assume e: "e > 0" then obtain N where N: "\n\N. \x. x \ s \ cmod (f' n x - g' x) \ e" - by (metis conv) + by (metis conv) have "\N. \n\N. \x\s. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" proof (rule exI [of _ N], clarify) fix n y h @@ -863,7 +863,7 @@ fixes s :: "complex set" assumes cvs: "convex s" and df: "\n x. x \ s \ (f n has_field_derivative f' n x) (at x within s)" - and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ s + and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ s \ cmod ((\i e" and "\x l. x \ s \ ((\n. f n x) sums l)" shows "\g. \x \ s. ((\n. f n x) sums g x) \ ((g has_field_derivative g' x) (at x within s))" @@ -871,9 +871,9 @@ from assms obtain x l where x: "x \ s" and sf: "((\n. f n x) sums l)" by blast { fix e::real assume e: "e > 0" - then obtain N where N: "\n x. n \ N \ x \ s + then obtain N where N: "\n x. n \ N \ x \ s \ cmod ((\i e" - by (metis conv) + by (metis conv) have "\N. \n\N. \x\s. \h. cmod ((\i e * cmod h" proof (rule exI [of _ N], clarify) fix n y h @@ -979,7 +979,7 @@ \ DERIV g (f x) :> inverse (f')" unfolding has_field_derivative_def apply (rule has_derivative_inverse_strong [of s x f g ]) - using assms + using assms by auto lemma has_complex_derivative_inverse_strong_x: @@ -993,7 +993,7 @@ \ DERIV g y :> inverse (f')" unfolding has_field_derivative_def apply (rule has_derivative_inverse_strong_x [of s g y f]) - using assms + using assms by auto subsection \Taylor on Complex Numbers\ @@ -1004,7 +1004,7 @@ by (induct n) auto lemma complex_taylor: - assumes s: "convex s" + assumes s: "convex s" and f: "\i x. x \ s \ i \ n \ (f i has_field_derivative f (Suc i) x) (at x within s)" and B: "\x. x \ s \ cmod (f (Suc n) x) \ B" and w: "w \ s" @@ -1019,14 +1019,14 @@ then have "u \ s" by (metis wzs subsetD) have "(\i\n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) + - f (Suc i) u * (z-u)^i / (fact i)) = + f (Suc i) u * (z-u)^i / (fact i)) = f (Suc n) u * (z-u) ^ n / (fact n)" proof (induction n) case 0 show ?case by simp next case (Suc n) have "(\i\Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) + - f (Suc i) u * (z-u) ^ i / (fact i)) = + f (Suc i) u * (z-u) ^ i / (fact i)) = f (Suc n) u * (z-u) ^ n / (fact n) + f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) - f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))" @@ -1056,7 +1056,7 @@ qed finally show ?case . qed - then have "((\v. (\i\n. f i v * (z - v)^i / (fact i))) + then have "((\v. (\i\n. f i v * (z - v)^i / (fact i))) has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n)) (at u within s)" apply (intro derivative_eq_intros) @@ -1081,19 +1081,19 @@ by simp also have "\ = f 0 z / (fact 0)" by (subst setsum_zero_power) simp - finally have "cmod (f 0 z - (\i\n. f i w * (z - w) ^ i / (fact i))) + finally have "cmod (f 0 z - (\i\n. f i w * (z - w) ^ i / (fact i))) \ cmod ((\i\n. f i w * (z - w) ^ i / (fact i)) - (\i\n. f i z * (z - z) ^ i / (fact i)))" by (simp add: norm_minus_commute) also have "... \ B * cmod (z - w) ^ n / (fact n) * cmod (w - z)" - apply (rule complex_differentiable_bound + apply (rule complex_differentiable_bound [where f' = "\w. f (Suc n) w * (z - w)^n / (fact n)" and s = "closed_segment w z", OF convex_closed_segment]) - apply (auto simp: ends_in_segment real_of_nat_def DERIV_subset [OF sum_deriv wzs] + apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs] norm_divide norm_mult norm_power divide_le_cancel cmod_bound) done also have "... \ B * cmod (z - w) ^ Suc n / (fact n)" - by (simp add: algebra_simps norm_minus_commute real_of_nat_def) + by (simp add: algebra_simps norm_minus_commute) finally show ?thesis . qed @@ -1142,7 +1142,7 @@ (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + (\i = 0..n. - f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - + f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - f (Suc i) u * (z-u) ^ i / (fact i))" by (simp only: diff_divide_distrib fact_cancel ac_simps) also have "... = f (Suc 0) u - @@ -1152,7 +1152,7 @@ by (subst setsum_Suc_diff) auto also have "... = f (Suc n) u * (z-u) ^ n / (fact n)" by (simp only: algebra_simps diff_divide_distrib fact_cancel) - finally have "(\i = 0..n. (f (Suc i) u * (z - u) ^ i + finally have "(\i = 0..n. (f (Suc i) u * (z - u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = f (Suc n) u * (z - u) ^ n / (fact n)" . then have "((\u. \i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Complex_Transcendental.thy --- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy Tue Nov 10 14:18:41 2015 +0000 @@ -192,8 +192,8 @@ lemma exp_eq_1: "exp z = 1 \ Re(z) = 0 \ (\n::int. Im(z) = of_int (2 * n) * pi)" apply auto apply (metis exp_eq_one_iff norm_exp_eq_Re norm_one) -apply (metis Re_exp cos_one_2pi_int mult.commute mult.left_neutral norm_exp_eq_Re norm_one one_complex.simps(1) real_of_int_def) -by (metis Im_exp Re_exp complex_Re_Im_cancel_iff cos_one_2pi_int sin_double Re_complex_of_real complex_Re_numeral exp_zero mult.assoc mult.left_commute mult_eq_0_iff mult_numeral_1 numeral_One of_real_0 real_of_int_def sin_zero_iff_int2) +apply (metis Re_exp cos_one_2pi_int mult.commute mult.left_neutral norm_exp_eq_Re norm_one one_complex.simps(1)) +by (metis Im_exp Re_exp complex_Re_Im_cancel_iff cos_one_2pi_int sin_double Re_complex_of_real complex_Re_numeral exp_zero mult.assoc mult.left_commute mult_eq_0_iff mult_numeral_1 numeral_One of_real_0 sin_zero_iff_int2) lemma exp_eq: "exp w = exp z \ (\n::int. w = z + (of_int (2 * n) * pi) * ii)" (is "?lhs = ?rhs") @@ -227,7 +227,7 @@ { assume "sin y = sin x" "cos y = cos x" then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp - then have "\n::int. y-x = real n * 2 * pi" + then have "\n::int. y-x = n * 2 * pi" using cos_one_2pi_int by blast } then show ?thesis apply (auto simp: sin_add cos_add) @@ -519,8 +519,8 @@ have *: "cmod (sin z - (\i\n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" - proof (rule complex_taylor [of "closed_segment 0 z" n - "\k x. (-1)^(k div 2) * (if even k then sin x else cos x)" + proof (rule complex_taylor [of "closed_segment 0 z" n + "\k x. (-1)^(k div 2) * (if even k then sin x else cos x)" "exp\Im z\" 0 z, simplified]) fix k x show "((\x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative @@ -622,7 +622,7 @@ then have "sin t' = sin t \ cos t' = cos t" apply (simp add: exp_Euler sin_of_real cos_of_real) by (metis Complex_eq complex.sel) - then obtain n::int where n: "t' = t + 2 * real n * pi" + then obtain n::int where n: "t' = t + 2 * n * pi" by (auto simp: sin_cos_eq_iff) then have "n=0" apply (rule_tac z=n in int_cases) @@ -752,7 +752,7 @@ by blast qed -corollary Arg_gt_0: +corollary Arg_gt_0: assumes "z \ \ \ Re z < 0" shows "Arg z > 0" using Arg_eq_0 Arg_ge_0 assms dual_order.strict_iff_order by fastforce @@ -963,7 +963,7 @@ lemma Re_Ln [simp]: "z \ 0 \ Re(Ln z) = ln(norm z)" by (metis exp_Ln assms ln_exp norm_exp_eq_Re) -corollary ln_cmod_le: +corollary ln_cmod_le: assumes z: "z \ 0" shows "ln (cmod z) \ cmod (Ln z)" using norm_exp [of "Ln z", simplified exp_Ln [OF z]] @@ -1205,7 +1205,7 @@ corollary Ln_divide_of_real: "\r > 0; z \ 0\ \ Ln(z / of_real r) = Ln(z) - ln r" using Ln_times_of_real [of "inverse r" z] -by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric] +by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric] del: of_real_inverse) lemma Ln_minus: @@ -1264,7 +1264,7 @@ lemma Ln_of_nat: "0 < n \ Ln (of_nat n) = of_real (ln (of_nat n))" by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all -lemma Ln_of_nat_over_of_nat: +lemma Ln_of_nat_over_of_nat: assumes "m > 0" "n > 0" shows "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))" proof - @@ -1279,7 +1279,7 @@ subsection\Relation between Ln and Arg, and hence continuity of Arg\ -lemma Arg_Ln: +lemma Arg_Ln: assumes "0 < Arg z" shows "Arg z = Im(Ln(-z)) + pi" proof (cases "z = 0") case True @@ -1307,7 +1307,7 @@ finally show ?thesis . qed -lemma continuous_at_Arg: +lemma continuous_at_Arg: assumes "z \ \ \ Re z < 0" shows "continuous (at z) Arg" proof - @@ -1322,7 +1322,7 @@ qed text\Relation between Arg and arctangent in upper halfplane\ -lemma Arg_arctan_upperhalf: +lemma Arg_arctan_upperhalf: assumes "0 < Im z" shows "Arg z = pi/2 - arctan(Re z / Im z)" proof (cases "z = 0") @@ -1340,24 +1340,24 @@ done qed -lemma Arg_eq_Im_Ln: - assumes "0 \ Im z" "0 < Re z" +lemma Arg_eq_Im_Ln: + assumes "0 \ Im z" "0 < Re z" shows "Arg z = Im (Ln z)" proof (cases "z = 0 \ Im z = 0") case True then show ?thesis - using assms Arg_eq_0 complex_is_Real_iff + using assms Arg_eq_0 complex_is_Real_iff apply auto by (metis Arg_eq_0_pi Arg_eq_pi Im_Ln_eq_0 Im_Ln_eq_pi less_numeral_extra(3) zero_complex.simps(1)) next - case False + case False then have "Arg z > 0" using Arg_gt_0 complex_is_Real_iff by blast then show ?thesis - using assms False + using assms False by (subst Arg_Ln) (auto simp: Ln_minus) qed -lemma continuous_within_upperhalf_Arg: +lemma continuous_within_upperhalf_Arg: assumes "z \ 0" shows "continuous (at z within {z. 0 \ Im z}) Arg" proof (cases "z \ \ & 0 \ Re z") @@ -1369,7 +1369,7 @@ using assms by (auto simp: complex_is_Real_iff complex_neq_0) then have [simp]: "Arg z = 0" "Im (Ln z) = 0" by (auto simp: Arg_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff) - show ?thesis + show ?thesis proof (clarsimp simp add: continuous_within Lim_within dist_norm) fix e::real assume "0 < e" @@ -1397,12 +1397,12 @@ apply (auto simp: continuous_on_eq_continuous_within) by (metis Diff_subset continuous_within_subset continuous_within_upperhalf_Arg) -lemma open_Arg_less_Int: +lemma open_Arg_less_Int: assumes "0 \ s" "t \ 2*pi" shows "open ({y. s < Arg y} \ {y. Arg y < t})" proof - have 1: "continuous_on (UNIV - {z \ \. 0 \ Re z}) Arg" - using continuous_at_Arg continuous_at_imp_continuous_within + using continuous_at_Arg continuous_at_imp_continuous_within by (auto simp: continuous_on_eq_continuous_within set_diff_eq) have 2: "open (UNIV - {z \ \. 0 \ Re z})" by (simp add: closed_Collect_conj closed_complex_Reals closed_halfspace_Re_ge open_Diff) @@ -1413,7 +1413,7 @@ using assms by (auto simp: Arg_real) ultimately show ?thesis - using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \ {z. Re z < t}"] + using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \ {z. Re z < t}"] by auto qed @@ -1521,11 +1521,11 @@ thus "eventually (\z. z powr r = Exp (r * Ln z)) (nhds z)" unfolding powr_def by eventually_elim simp - have "((\z. Exp (r * Ln z)) has_field_derivative Exp (r * Ln z) * (inverse z * r)) (at z)" + have "((\z. Exp (r * Ln z)) has_field_derivative Exp (r * Ln z) * (inverse z * r)) (at z)" using assms by (auto intro!: derivative_eq_intros has_field_derivative_powr) also have "Exp (r * Ln z) * (inverse z * r) = r * z powr (r - 1)" unfolding powr_def by (simp add: assms exp_diff field_simps) - finally show "((\z. Exp (r * Ln z)) has_field_derivative r * z powr (r - 1)) (at z)" + finally show "((\z. Exp (r * Ln z)) has_field_derivative r * z powr (r - 1)) (at z)" by simp qed @@ -1553,13 +1553,13 @@ subsection\Some Limits involving Logarithms\ - + lemma lim_Ln_over_power: fixes s::complex assumes "0 < Re s" shows "((\n. Ln n / (n powr s)) ---> 0) sequentially" proof (simp add: lim_sequentially dist_norm, clarify) - fix e::real + fix e::real assume e: "0 < e" have "\xo>0. \x\xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe) @@ -1591,7 +1591,7 @@ apply (rule_tac x="nat (ceiling (exp xo))" in exI) apply clarify apply (drule_tac x="ln n" in spec) - apply (auto simp: real_of_nat_def exp_less_mono nat_ceiling_le_eq not_le) + apply (auto simp: exp_less_mono nat_ceiling_le_eq not_le) apply (metis exp_less_mono exp_ln not_le of_nat_0_less_iff) done qed @@ -1613,13 +1613,13 @@ using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm - Ln_Reals_eq norm_powr_real_powr norm_divide real_of_nat_def) + Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_ln_over_n: "((\n. ln(real_of_nat n) / of_nat n) ---> 0) sequentially" using lim_ln_over_power [of 1, THEN filterlim_sequentially_Suc [THEN iffD2]] apply (subst filterlim_sequentially_Suc [symmetric]) - apply (simp add: lim_sequentially dist_norm real_of_nat_def) + apply (simp add: lim_sequentially dist_norm) done lemma lim_1_over_complex_power: @@ -1645,7 +1645,7 @@ using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) - apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide real_of_nat_def) + apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_1_over_Ln: "((\n. 1 / Ln(of_nat n)) ---> 0) sequentially" @@ -1677,7 +1677,7 @@ using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) - apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide real_of_nat_def) + apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) done @@ -1842,7 +1842,7 @@ also have "... \ False" using assms ge_pi2 apply (auto simp: algebra_simps) - by (metis abs_mult_pos not_less not_real_of_nat_less_zero real_of_nat_numeral) + by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral) finally have *: "exp (\*z)*exp (\*z) + 1 \ 0" by (auto simp: add.commute minus_unique) show ?thesis @@ -1939,7 +1939,7 @@ apply (subst exp_Ln, auto) apply (simp_all add: cmod_def complex_eq_iff) apply (auto simp: divide_simps) - apply (metis power_one realpow_two_sum_zero_iff zero_neq_one, algebra) + apply (metis power_one sum_power2_eq_zero_iff zero_neq_one, algebra) done lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Nov 10 14:18:41 2015 +0000 @@ -4596,7 +4596,7 @@ using separating_hyperplane_closed_0[OF convex_convex_hull, of c] using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) using subset_hull[of convex, OF assms(1), symmetric, of c] - by auto + by force then have "\x. norm x = 1 \ (\y\c. 0 \ inner y x)" apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) using hull_subset[of c convex] @@ -4624,7 +4624,7 @@ proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] obtain a where "a \ 0" "\x\{x - y |x y. x \ t \ y \ s}. 0 \ inner a x" - using assms(3-5) by auto + using assms(3-5) by fastforce then have *: "\x y. x \ t \ y \ s \ inner a y \ inner a x" by (force simp add: inner_diff) then have bdd: "bdd_above ((op \ a)`s)" @@ -4832,11 +4832,10 @@ have *: "\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto show ?thesis - unfolding real_add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)] + unfolding add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)] and setsum.distrib[symmetric] and * using assms(2) - apply assumption - done + by assumption qed lemma radon_v_lemma: @@ -5757,7 +5756,7 @@ subsection \Another intermediate value theorem formulation\ lemma ivt_increasing_component_on_1: - fixes f :: "real \ 'a::euclidean_space" + fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f a)\k \ y" "y \ (f b)\k" @@ -6204,7 +6203,7 @@ apply (erule (1) order_trans [OF Basis_le_norm]) done have e': "e = (\(i::'a)\Basis. d)" - by (simp add: d_def real_of_nat_def DIM_positive) + by (simp add: d_def DIM_positive) show "convex hull c \ cball x e" unfolding 2 apply clarsimp @@ -6326,7 +6325,7 @@ apply (simp add: scaleR_2) done show ?t3 - unfolding midpoint_def dist_norm + unfolding midpoint_def dist_norm apply (rule *) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) @@ -6910,7 +6909,7 @@ then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" - unfolding setsum.distrib setsum_constant real_eq_of_nat + unfolding setsum.distrib setsum_constant by (auto simp add: Suc_le_eq) finally show "(\i\Basis. 0 \ y \ i) \ setsum (op \ y) Basis \ 1" proof safe @@ -6957,7 +6956,7 @@ apply auto done also have "\ < 1" - unfolding setsum_constant real_eq_of_nat divide_inverse[symmetric] + unfolding setsum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "setsum (op \ ?a) ?D < 1" by auto qed @@ -7094,8 +7093,7 @@ then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" - unfolding setsum.distrib setsum_constant real_eq_of_nat - using \0 < card d\ + unfolding setsum.distrib setsum_constant using \0 < card d\ by auto finally show "setsum (op \ y) d \ 1" . @@ -7164,7 +7162,7 @@ have "setsum (op \ ?a) ?D = setsum (\i. inverse (2 * real (card d))) ?D" by (rule setsum.cong) (rule refl, rule **) also have "\ < 1" - unfolding setsum_constant real_eq_of_nat divide_real_def[symmetric] + unfolding setsum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "setsum (op \ ?a) ?D < 1" by auto next @@ -9713,7 +9711,7 @@ unfolding setdist_def by (auto intro!: bdd_belowI [where m=0] cInf_lower) -lemma le_setdist_iff: +lemma le_setdist_iff: "d \ setdist s t \ (\x \ s. \y \ t. d \ dist x y) \ (s = {} \ t = {} \ d \ 0)" apply (cases "s = {} \ t = {}") @@ -9723,7 +9721,7 @@ apply (auto simp: intro: le_setdistI) done -lemma setdist_ltE: +lemma setdist_ltE: assumes "setdist s t < b" "s \ {}" "t \ {}" obtains x y where "x \ s" "y \ t" "dist x y < b" using assms @@ -9745,7 +9743,7 @@ using setdist_pos_le by fastforce next case False - have "\x. x \ s \ setdist s t - dist x a \ setdist {a} t" + have "\x. x \ s \ setdist s t - dist x a \ setdist {a} t" apply (rule le_setdistI, blast) using False apply (fastforce intro: le_setdistI) apply (simp add: algebra_simps) @@ -9840,7 +9838,7 @@ assumes s: "compact s" and t: "closed t" shows "setdist s t = 0 \ s = {} \ t = {} \ s \ t \ {}" apply (cases "s = {} \ t = {}", force) - using setdist_compact_closed [OF s t] + using setdist_compact_closed [OF s t] apply (force intro: setdist_eq_0I ) done @@ -9863,29 +9861,29 @@ assumes "bounded s \ bounded t" shows "setdist s t = 0 \ s = {} \ t = {} \ closure s \ closure t \ {}" apply (cases "s = {} \ t = {}", force) - using setdist_eq_0_compact_closed [of "closure s" "closure t"] - setdist_eq_0_closed_compact [of "closure s" "closure t"] assms + using setdist_eq_0_compact_closed [of "closure s" "closure t"] + setdist_eq_0_closed_compact [of "closure s" "closure t"] assms apply (force simp add: bounded_closure compact_eq_bounded_closed) done -lemma setdist_unique: +lemma setdist_unique: "\a \ s; b \ t; \x y. x \ s \ y \ t ==> dist a b \ dist x y\ \ setdist s t = dist a b" by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) -lemma setdist_closest_point: +lemma setdist_closest_point: "\closed s; s \ {}\ \ setdist {a} s = dist a (closest_point s a)" apply (rule setdist_unique) using closest_point_le apply (auto simp: closest_point_in_set) done -lemma setdist_eq_0_sing_1 [simp]: +lemma setdist_eq_0_sing_1 [simp]: fixes s :: "'a::euclidean_space set" shows "setdist {x} s = 0 \ s = {} \ x \ closure s" by (auto simp: setdist_eq_0_bounded) -lemma setdist_eq_0_sing_2 [simp]: +lemma setdist_eq_0_sing_2 [simp]: fixes s :: "'a::euclidean_space set" shows "setdist s {x} = 0 \ s = {} \ x \ closure s" by (auto simp: setdist_eq_0_bounded) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Integration.thy --- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Nov 10 14:18:41 2015 +0000 @@ -2860,7 +2860,7 @@ show "norm ((\(x, k)\p m. content k *\<^sub>R f x) - (\(x, k)\p n. content k *\<^sub>R f x)) < e" apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]]) apply(subst *) - using dp p(1) mn d(2) real_of_nat_def by auto + using dp p(1) mn d(2) by auto qed qed then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN LIMSEQ_D] @@ -4518,7 +4518,7 @@ using True by (auto simp add: field_simps) then obtain M :: nat where M: "M \ 0" "0 < inverse (real_of_nat M)" "inverse (of_nat M) < e / 4 / content (cbox a b)" - by (subst (asm) real_arch_inv) (auto simp: real_of_nat_def) + by (subst (asm) real_arch_inv) auto show "\M. \m\M. \n\M. dist (i m) (i n) < e" proof (rule exI [where x=M], clarify) fix m n @@ -6091,8 +6091,7 @@ have Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})" unfolding Dg_def - by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def - fact_eq real_eq_of_nat[symmetric] divide_simps) + by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps) let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, OF \p > 0\ g0 Dg f0 Df] @@ -6420,17 +6419,17 @@ shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" proof - let ?I = "\a b. integral {a..b} f" - { fix e::real + { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b}; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" - if y: "y \ {a..b}" and yx: "\y - x\ < d" for y + if y: "y \ {a..b}" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) - then have Idiff: "?I a y - ?I a x = ?I x y" + then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y - x) *\<^sub>R f x) {x..y}" apply (rule has_integral_sub) @@ -6448,7 +6447,7 @@ case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) - then have Idiff: "?I a x - ?I a y = ?I y x" + then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" apply (rule has_integral_sub) @@ -6466,8 +6465,8 @@ by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b}. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \ e * \y - x\" - using \d>0\ by blast - } + using \d>0\ by blast + } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed @@ -9485,8 +9484,7 @@ defer unfolding divide_inverse setsum_left_distrib[symmetric] unfolding divide_inverse[symmetric] - using * - apply (auto simp add: field_simps real_eq_of_nat) + using * apply (auto simp add: field_simps) done next case 2 diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Linear_Algebra.thy --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy Tue Nov 10 14:18:41 2015 +0000 @@ -570,11 +570,11 @@ subsection \Archimedean properties and useful consequences\ lemma real_arch_simple: "\n::nat. x \ real n" - unfolding real_of_nat_def by (rule ex_le_of_nat) + by (rule ex_le_of_nat) lemma real_arch_inv: "0 < e \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" using reals_Archimedean[of e] less_trans[of 0 "1 / real n" e for n::nat] - by (auto simp add: field_simps cong: conj_cong simp del: real_of_nat_Suc) + by (auto simp add: field_simps cong: conj_cong simp del: of_nat_Suc) text\Bernoulli's inequality\ lemma Bernoulli_inequality: @@ -589,7 +589,7 @@ have "1 + Suc n * x \ 1 + (Suc n)*x + n * x^2" by (simp add: algebra_simps) also have "... = (1 + x) * (1 + n*x)" - by (auto simp: power2_eq_square algebra_simps real_of_nat_Suc) + by (auto simp: power2_eq_square algebra_simps of_nat_Suc) also have "... \ (1 + x) ^ Suc n" using Suc.hyps assms mult_left_mono by fastforce finally show ?case . @@ -668,7 +668,7 @@ (\n. P (inverse (real (Suc n)))) \ 0 < e \ P e" apply (rule forall_pos_mono) apply auto - apply (metis Suc_pred real_of_nat_Suc) + apply (metis Suc_pred of_nat_Suc) done lemma real_archimedian_rdiv_eq_0: @@ -1461,8 +1461,7 @@ by (intro add_mono norm_bound_Basis_le i fPs) auto finally show "(\x\P. \f x \ i\) \ 2*e" by simp qed - also have "\ = 2 * real DIM('n) * e" - by (simp add: real_of_nat_def) + also have "\ = 2 * real DIM('n) * e" by simp finally show ?thesis . qed @@ -1629,12 +1628,12 @@ shows "independent S \ finite S \ card S \ DIM('a)" using independent_span_bound[OF finite_Basis, of S] by auto -corollary +corollary fixes S :: "'a::euclidean_space set" assumes "independent S" shows independent_imp_finite: "finite S" and independent_card_le:"card S \ DIM('a)" using assms independent_bound by auto - + lemma dependent_biggerset: fixes S :: "'a::euclidean_space set" shows "(finite S \ card S > DIM('a)) \ dependent S" @@ -2856,7 +2855,6 @@ by (simp add: zero_le_mult_iff infnorm_pos_le) have th1: "x \ x \ (sqrt (real ?d) * infnorm x)\<^sup>2" unfolding power_mult_distrib d2 - unfolding real_of_nat_def apply (subst euclidean_inner) apply (subst power2_abs[symmetric]) apply (rule order_trans[OF setsum_bounded_above[where K="\infnorm x\\<^sup>2"]]) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/PolyRoots.thy --- a/src/HOL/Multivariate_Analysis/PolyRoots.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/PolyRoots.thy Tue Nov 10 14:18:41 2015 +0000 @@ -20,7 +20,7 @@ fixes x :: "'a::{comm_ring,division_ring}" shows "(\i\n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))" using setsum_gp_basic[of x n] - by (simp add: real_of_nat_def mult.commute divide_simps) + by (simp add: mult.commute divide_simps) lemma setsum_power_add: fixes x :: "'a::{comm_ring,monoid_mult}" @@ -60,8 +60,8 @@ (if n < m then 0 else if x = 1 then of_nat((n + 1) - m) else (x^m - x^Suc n) / (1 - x))" -using setsum_gp_multiplied [of m n x] -apply (auto simp: real_of_nat_def) +using setsum_gp_multiplied [of m n x] +apply auto by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq) lemma setsum_gp_offset: @@ -75,15 +75,15 @@ fixes x :: "'a::{comm_ring,division_ring}" shows "(\iBasics about polynomial functions: extremal behaviour and root counts.\ lemma sub_polyfun: fixes x :: "'a::{comm_ring,monoid_mult}" - shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = + shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\jk= Suc j..n. a k * y^(k - Suc j) * x^j)" proof - - have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = + have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: algebra_simps setsum_subtractf [symmetric]) also have "... = (\i\n. a i * (x - y) * (\ji\n. a i * x^i) - (\i\n. a i * y^i) = + shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\jkb. \z. (\i\n. c i * z^i) = + shows "\b. \z. (\i\n. c i * z^i) = (z-a) * (\ii\n. c i * a^i)" proof - { fix z - have "(\i\n. c i * z^i) - (\i\n. c i * a^i) = + have "(\i\n. c i * z^i) - (\i\n. c i * a^i) = (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j)" by (simp add: sub_polyfun setsum_left_distrib) - then have "(\i\n. c i * z^i) = + then have "(\i\n. c i * z^i) = (z - a) * (\jk = Suc j..n. c k * a^(k - Suc j)) * z^j) + (\i\n. c i * a^i)" by (simp add: algebra_simps) } then show ?thesis - by (intro exI allI) + by (intro exI allI) qed lemma polyfun_linear_factor_root: @@ -141,7 +141,7 @@ shows "\M. \z. M \ norm z \ norm(\i\n. c i * z^i) \ e * norm(z) ^ Suc n" proof (induction n) case 0 - show ?case + show ?case by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms) next case (Suc n) @@ -153,14 +153,14 @@ then have norm1: "0 < norm z" "M \ norm z" "(e + norm (c (Suc n))) / e \ norm z" by auto then have norm2: "(e + norm (c (Suc n))) \ e * norm z" "(norm z * norm z ^ n) > 0" - apply (metis assms less_divide_eq mult.commute not_le) + apply (metis assms less_divide_eq mult.commute not_le) using norm1 apply (metis mult_pos_pos zero_less_power) done have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) = (e + norm (c (Suc n))) * (norm z * norm z ^ n)" by (simp add: norm_mult norm_power algebra_simps) also have "... \ (e * norm z) * (norm z * norm z ^ n)" - using norm2 by (metis real_mult_le_cancel_iff1) + using norm2 by (metis real_mult_le_cancel_iff1) also have "... = e * (norm z * (norm z * norm z ^ n))" by (simp add: algebra_simps) finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) @@ -171,7 +171,7 @@ qed lemma norm_lemma_xy: "\abs b + 1 \ norm(y) - a; norm(x) \ a\ \ b \ norm(x + y)" - by (metis abs_add_one_not_less_self add.commute diff_le_eq dual_order.trans le_less_linear + by (metis abs_add_one_not_less_self add.commute diff_le_eq dual_order.trans le_less_linear norm_diff_ineq) lemma polyfun_extremal: @@ -192,7 +192,7 @@ next case False with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n] - obtain M where M: "\z. M \ norm z \ + obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ norm (c (Suc n)) / 2 * norm z ^ Suc n" by auto show ?thesis @@ -202,7 +202,7 @@ assume les: "M \ norm z" "1 \ norm z" "(\B\ * 2 + 2) / norm (c (Suc n)) \ norm z" then have "\B\ * 2 + 2 \ norm z * norm (c (Suc n))" by (metis False pos_divide_le_eq zero_less_norm_iff) - then have "\B\ * 2 + 2 \ norm z ^ (Suc n) * norm (c (Suc n))" + then have "\B\ * 2 + 2 \ norm z ^ (Suc n) * norm (c (Suc n))" by (metis \1 \ norm z\ order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc) then show "B \ norm ((\i\n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les apply auto @@ -235,7 +235,7 @@ by simp then have extr_prem: "~ (\k\n. b k \ 0) \ \k. k \ 0 \ k \ Suc n \ c k \ 0" by (metis Suc.prems le0 minus_zero mult_zero_right) - have "\k\n. b k \ 0" + have "\k\n. b k \ 0" apply (rule ccontr) using polyfun_extremal [OF extr_prem, of 1] apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc) @@ -257,10 +257,10 @@ fixes c :: "nat \ 'a::{comm_ring,real_normed_div_algebra}" shows "finite {z. (\i\n. c i * z^i) = 0} \ (\k. k \ n \ c k \ 0)" proof (cases " \k\n. c k \ 0") - case True then show ?thesis + case True then show ?thesis by (blast intro: polyfun_rootbound_finite) next - case False then show ?thesis + case False then show ?thesis by (auto simp: infinite_UNIV_char_0) qed diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Nov 10 14:18:41 2015 +0000 @@ -707,14 +707,14 @@ from 1 2 show ?lhs unfolding openin_open open_dist by fast qed - + lemma connected_open_in: "connected s \ ~(\e1 e2. openin (subtopology euclidean s) e1 \ openin (subtopology euclidean s) e2 \ s \ e1 \ e2 \ e1 \ e2 = {} \ e1 \ {} \ e2 \ {})" apply (simp add: connected_def openin_open, safe) - apply (simp_all, blast+) + apply (simp_all, blast+) --\slow\ done lemma connected_open_in_eq: @@ -768,7 +768,7 @@ apply (simp add: connected_closed_in, safe) apply blast by (metis Int_lower1 Un_subset_iff closedin_closed subset_antisym) - + text \These "transitivity" results are handy too\ lemma openin_trans[trans]: @@ -1002,7 +1002,7 @@ by (simp add: power_divide) qed auto also have "\ = e" - using \0 < e\ by (simp add: real_eq_of_nat) + using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) qed (insert a b, auto simp: box_def) @@ -1248,14 +1248,14 @@ lemma UN_box_eq_UNIV: "(\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" proof - - have "\x \ b\ < real (ceiling (Max ((\b. \x \ b\)`Basis)) + 1)" + have "\x \ b\ < real_of_int (ceiling (Max ((\b. \x \ b\)`Basis)) + 1)" if [simp]: "b \ Basis" for x b :: 'a proof - - have "\x \ b\ \ real (ceiling \x \ b\)" - by (rule real_of_int_ceiling_ge) - also have "\ \ real (ceiling (Max ((\b. \x \ b\)`Basis)))" + have "\x \ b\ \ real_of_int (ceiling \x \ b\)" + by (rule le_of_int_ceiling) + also have "\ \ real_of_int (ceiling (Max ((\b. \x \ b\)`Basis)))" by (auto intro!: ceiling_mono) - also have "\ < real (ceiling (Max ((\b. \x \ b\)`Basis)) + 1)" + also have "\ < real_of_int (ceiling (Max ((\b. \x \ b\)`Basis)) + 1)" by simp finally show ?thesis . qed @@ -1277,15 +1277,11 @@ unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff by (meson order_trans le_less_trans less_le_trans less_trans)+ -lemma is_interval_empty: - "is_interval {}" - unfolding is_interval_def - by simp - -lemma is_interval_univ: - "is_interval UNIV" - unfolding is_interval_def - by simp +lemma is_interval_empty [iff]: "is_interval {}" + unfolding is_interval_def by simp + +lemma is_interval_univ [iff]: "is_interval UNIV" + unfolding is_interval_def by simp lemma mem_is_intervalI: assumes "is_interval s" @@ -2162,7 +2158,7 @@ lemma closed_in_connected_component: "closedin (subtopology euclidean s) (connected_component_set s x)" proof (cases "connected_component_set s x = {}") - case True then show ?thesis + case True then show ?thesis by (metis closedin_empty) next case False @@ -4383,7 +4379,7 @@ obtain r where r: "subseq r" "monoseq (f \ r)" unfolding comp_def by (metis seq_monosub) then have "Bseq (f \ r)" - unfolding Bseq_eq_bounded using f by (auto intro: bounded_subset) + unfolding Bseq_eq_bounded using f by (force intro: bounded_subset) with r show "\l r. subseq r \ (f \ r) ----> l" using Bseq_monoseq_convergent[of "f \ r"] by (auto simp: convergent_def) qed @@ -4579,10 +4575,10 @@ assumes "0 < e" obtains n :: nat where "1 / (Suc n) < e" proof atomize_elim - have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))" + have "1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))" by (rule divide_strict_left_mono) (auto simp: \0 < e\) also have "1 / (ceiling (1/e)) \ 1 / (1/e)" - by (rule divide_left_mono) (auto simp: \0 < e\) + by (rule divide_left_mono) (auto simp: \0 < e\ ceiling_correct) also have "\ = e" by simp finally show "\n. 1 / real (Suc n) < e" .. qed @@ -4665,7 +4661,7 @@ fix r :: real and N n m assume "1 / Suc N < r" "Suc N \ n" "Suc N \ m" then have "(f \ t) n \ F (Suc N)" "(f \ t) m \ F (Suc N)" "2 * e N < r" - using F_dec t by (auto simp: e_def field_simps real_of_nat_Suc) + using F_dec t by (auto simp: e_def field_simps of_nat_Suc) with F[of N] obtain x where "dist x ((f \ t) n) < e N" "dist x ((f \ t) m) < e N" by (auto simp: subset_eq) with dist_triangle[of "(f \ t) m" "(f \ t) n" x] \2 * e N < r\ @@ -5130,7 +5126,7 @@ unfolding continuous_within Lim_within ball_def subset_eq apply (auto simp add: dist_commute) apply (erule_tac x=e in allE) - apply auto + apply auto done qed @@ -5338,7 +5334,7 @@ fix n :: nat assume "n \ N" then have "inverse (real n + 1) < inverse (real N)" - using real_of_nat_ge_zero and \N\0\ by auto + using of_nat_0_le_iff and \N\0\ by auto also have "\ < e" using N by auto finally have "inverse (real n + 1) < e" by auto then have "dist (x n) (y n) < e" @@ -5946,15 +5942,12 @@ lemma bilinear_continuous_within_compose: "continuous (at x within s) f \ continuous (at x within s) g \ bounded_bilinear h \ continuous (at x within s) (\x. h (f x) (g x))" - unfolding continuous_within - using Lim_bilinear[of f "f x"] - by auto + by (rule Limits.bounded_bilinear.continuous) lemma bilinear_continuous_on_compose: "continuous_on s f \ continuous_on s g \ bounded_bilinear h \ continuous_on s (\x. h (f x) (g x))" - unfolding continuous_on_def - by (fast elim: bounded_bilinear.tendsto) + by (rule Limits.bounded_bilinear.continuous_on) text \Preservation of compactness and connectedness under continuous function.\ @@ -6024,7 +6017,7 @@ then show ?thesis unfolding connected_clopen by auto qed - + lemma connected_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "linear f" and "connected s" @@ -6773,7 +6766,7 @@ and x: "x \ closure(s)" and xlo: "\x. x \ s ==> f(x) \ a" shows "f(x) \ a" - using image_closure_subset [OF f] + using image_closure_subset [OF f] using image_closure_subset [OF f] closed_halfspace_le [of "1::real" a] assms by force @@ -7140,15 +7133,12 @@ then have "\N::nat. inverse (real (N + 1)) < e" using real_arch_inv[of e] apply (auto simp add: Suc_pred') - apply (metis Suc_pred' real_of_nat_Suc) + apply (metis Suc_pred' of_nat_Suc) done - then obtain N :: nat where "inverse (real (N + 1)) < e" + then obtain N :: nat where N: "inverse (real (N + 1)) < e" by auto - then have "\n\N. inverse (real n + 1) < e" - apply auto - apply (metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans - real_of_nat_Suc real_of_nat_Suc_gt_zero) - done + have "inverse (real n + 1) < e" if "N \ n" for n + by (auto intro!: that le_less_trans [OF _ N]) then have "\N::nat. \n\N. inverse (real n + 1) < e" by auto } then have "((\n. inverse (real n + 1)) ---> 0) sequentially" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Multivariate_Analysis/Weierstrass.thy --- a/src/HOL/Multivariate_Analysis/Weierstrass.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Multivariate_Analysis/Weierstrass.thy Tue Nov 10 14:18:41 2015 +0000 @@ -5,9 +5,6 @@ begin -(*FIXME: simplification changes to be enforced globally*) -declare of_nat_Suc [simp del] - (*Power.thy:*) declare power_divide [where b = "numeral w" for w, simp del] @@ -45,7 +42,7 @@ lemma sum_k_Bernstein [simp]: "(\k = 0..n. real k * Bernstein n k x) = of_nat n * x" apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric]) apply (simp add: setsum_left_distrib) - apply (auto simp: Bernstein_def real_of_nat_def algebra_simps realpow_num_eq_if intro!: setsum.cong) + apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong) done lemma sum_kk_Bernstein [simp]: "(\ k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2" @@ -54,7 +51,7 @@ apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric]) apply (simp add: setsum_left_distrib) apply (rule setsum.cong [OF refl]) - apply (simp add: Bernstein_def power2_eq_square algebra_simps real_of_nat_def) + apply (simp add: Bernstein_def power2_eq_square algebra_simps) apply (rename_tac k) apply (subgoal_tac "k = 0 \ k = 1 \ (\k'. k = Suc (Suc k'))") apply (force simp add: field_simps of_nat_Suc power2_eq_square) @@ -90,13 +87,13 @@ using e \0 by simp also have "... \ M * 4" using \0\M\ by simp - finally have [simp]: "real (nat \4 * M / (e * d\<^sup>2)\) = real \4 * M / (e * d\<^sup>2)\" + finally have [simp]: "real_of_int (nat \4 * M / (e * d\<^sup>2)\) = real_of_int \4 * M / (e * d\<^sup>2)\" using \0\M\ e \0 - by (simp add: Real.real_of_nat_Suc field_simps) + by (simp add: of_nat_Suc field_simps) have "4*M/(e*d\<^sup>2) + 1 \ real (Suc (nat\4*M/(e*d\<^sup>2)\))" - by (simp add: Real.real_of_nat_Suc) + by (simp add: of_nat_Suc real_nat_ceiling_ge) also have "... \ real n" - using n by (simp add: Real.real_of_nat_Suc field_simps) + using n by (simp add: of_nat_Suc field_simps) finally have nbig: "4*M/(e*d\<^sup>2) + 1 \ real n" . have sum_bern: "(\k = 0..n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n" proof - @@ -296,7 +293,7 @@ using pf01 by force moreover have "p x \ 1" using subU cardp t - apply (simp add: p_def divide_simps setsum_nonneg real_of_nat_def) + apply (simp add: p_def divide_simps setsum_nonneg) apply (rule setsum_bounded_above [where 'a=real and K=1, simplified]) using pf01 by force ultimately show ?thesis @@ -484,7 +481,7 @@ using ff by fastforce moreover have "pff x \ 1" using subA cardp t - apply (simp add: pff_def divide_simps setsum_nonneg real_of_nat_def) + apply (simp add: pff_def divide_simps setsum_nonneg) apply (rule setprod_mono [where g = "\x. 1", simplified]) using ff by fastforce ultimately show ?thesis @@ -571,7 +568,7 @@ def B \ "\j::nat. {x \ s. f x \ (j + 1/3)*e}" have ngt: "(n-1) * e \ normf f" "n\1" using e - apply (simp_all add: n_def field_simps real_of_nat_Suc) + apply (simp_all add: n_def field_simps of_nat_Suc) by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq) then have ge_fx: "(n-1) * e \ f x" if "x \ s" for x using f normf_upper that by fastforce @@ -606,7 +603,7 @@ have A0: "A 0 = {}" using fpos e by (fastforce simp: A_def) have An: "A n = s" - using e ngt f normf_upper by (fastforce simp: A_def field_simps real_of_nat_diff) + using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff) have Asub: "A j \ A i" if "i\j" for i j using e that apply (clarsimp simp: A_def) apply (erule order_trans, simp) @@ -631,7 +628,7 @@ then have Anj: "t \ A (j-1)" using Least_le by (fastforce simp add: j_def) then have fj2: "(j - 4/3)*e < f t" - using j1 t by (simp add: A_def real_of_nat_diff) + using j1 t by (simp add: A_def of_nat_diff) have ***: "xf i t \ e/n" if "i\j" for i using xfA [OF Ai] that by (simp add: less_eq_real_def) { fix i @@ -641,7 +638,7 @@ then have "t \ B i" using Anj e ge_fx [OF t] \1 \ n\ fpos [OF t] t apply (simp add: A_def B_def) - apply (clarsimp simp add: field_simps real_of_nat_diff not_le real_of_nat_Suc) + apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc) apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"]) apply auto done @@ -658,15 +655,15 @@ done also have "... \ e*j + e * ((Suc n - j)*e/n)" apply (rule add_mono) - apply (simp_all only: mult_le_cancel_left_pos e real_of_nat_def) + apply (simp_all only: mult_le_cancel_left_pos e) apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified]) using setsum_bounded_above [of "{j..n}" "\i. xf i t", OF ***] apply simp done also have "... \ j*e + e*(n - j + 1)*e/n " - using \1 \ n\ e by (simp add: field_simps del: real_of_nat_Suc) + using \1 \ n\ e by (simp add: field_simps del: of_nat_Suc) also have "... \ j*e + e*e" - using \1 \ n\ e j1 by (simp add: field_simps del: real_of_nat_Suc) + using \1 \ n\ e j1 by (simp add: field_simps del: of_nat_Suc) also have "... < (j + 1/3)*e" using e by (auto simp: field_simps) finally have gj1: "g t < (j + 1 / 3) * e" . @@ -678,7 +675,7 @@ next case True then have "(j - 4/3)*e < (j-1)*e - e^2" - using e by (auto simp: real_of_nat_diff algebra_simps power2_eq_square) + using e by (auto simp: of_nat_diff algebra_simps power2_eq_square) also have "... < (j-1)*e - ((j - 1)/n) * e^2" using e True jn by (simp add: power2_eq_square field_simps) also have "... = e * (j-1) * (1 - e/n)" @@ -688,7 +685,7 @@ apply simp apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]]) using True - apply (simp_all add: real_of_nat_def of_nat_Suc of_nat_diff) + apply (simp_all add: of_nat_Suc of_nat_diff) done also have "... \ g t" using jn e @@ -739,7 +736,7 @@ have "\ real (Suc n) < inverse e" using \N \ n\ N using less_imp_inverse_less by force then have "1 / (1 + real n) \ e" - using e by (simp add: field_simps real_of_nat_Suc) + using e by (simp add: field_simps of_nat_Suc) then have "\f x - g x\ < e" using n(2) x by auto } note * = this @@ -922,7 +919,7 @@ show ?case apply (rule_tac x="\i. c" in exI) apply (rule_tac x=0 in exI) - apply (auto simp: mult_ac real_of_nat_Suc) + apply (auto simp: mult_ac of_nat_Suc) done case (add f1 f2) then obtain a1 n1 a2 n2 where diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/CLim.thy --- a/src/HOL/NSA/CLim.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/CLim.thy Tue Nov 10 14:18:41 2015 +0000 @@ -140,15 +140,14 @@ "DERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))" apply (induct n) apply (drule_tac [2] DERIV_ident [THEN DERIV_mult]) -apply (auto simp add: distrib_right real_of_nat_Suc) +apply (auto simp add: distrib_right of_nat_Suc) apply (case_tac "n") -apply (auto simp add: ac_simps add.commute) +apply (auto simp add: ac_simps) done text{*Nonstandard version*} -lemma NSCDERIV_pow: - "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))" -by (simp add: NSDERIV_DERIV_iff del: of_real_real_of_nat_eq) +lemma NSCDERIV_pow: "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))" + by (metis CDERIV_pow NSDERIV_DERIV_iff One_nat_def) text{*Can't relax the premise @{term "x \ 0"}: it isn't continuous at zero*} lemma NSCDERIV_inverse: diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/HSEQ.thy --- a/src/HOL/NSA/HSEQ.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/HSEQ.thy Tue Nov 10 14:18:41 2015 +0000 @@ -225,11 +225,11 @@ by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_zero) lemma NSLIMSEQ_inverse_real_of_nat: "(%n. inverse(real(Suc n))) ----NS> 0" -by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat del: real_of_nat_Suc) +by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat del: of_nat_Suc) lemma NSLIMSEQ_inverse_real_of_nat_add: "(%n. r + inverse(real(Suc n))) ----NS> r" -by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add del: real_of_nat_Suc) +by (simp add: LIMSEQ_NSLIMSEQ_iff [symmetric] LIMSEQ_inverse_real_of_nat_add del: of_nat_Suc) lemma NSLIMSEQ_inverse_real_of_nat_add_minus: "(%n. r + -inverse(real(Suc n))) ----NS> r" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/HSeries.thy --- a/src/HOL/NSA/HSeries.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/HSeries.thy Tue Nov 10 14:18:41 2015 +0000 @@ -2,8 +2,8 @@ Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge -Converted to Isar and polished by lcp -*) +Converted to Isar and polished by lcp +*) section{*Finite Summation and Infinite Series for Hyperreals*} @@ -13,7 +13,7 @@ definition sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal" where - "sumhr = + "sumhr = (%(M,N,f). starfun2 (%m n. setsum f {m.. m then 0 else sumhr(m,n,f) + ( *f* f) n)" unfolding sumhr_app by transfer simp @@ -73,14 +73,14 @@ text{* other general version also needed *} lemma sumhr_fun_hypnat_eq: - "(\r. m \ r & r < n --> f r = g r) --> - sumhr(hypnat_of_nat m, hypnat_of_nat n, f) = + "(\r. m \ r & r < n --> f r = g r) --> + sumhr(hypnat_of_nat m, hypnat_of_nat n, f) = sumhr(hypnat_of_nat m, hypnat_of_nat n, g)" unfolding sumhr_app by transfer simp lemma sumhr_const: "!!n. sumhr(0, n, %i. r) = hypreal_of_hypnat n * hypreal_of_real r" -unfolding sumhr_app by transfer (simp add: real_of_nat_def) +unfolding sumhr_app by transfer simp lemma sumhr_less_bounds_zero [simp]: "!!m n. n < m ==> sumhr(m,n,f) = 0" unfolding sumhr_app by transfer simp @@ -98,7 +98,7 @@ text{*Infinite sums are obtained by summing to some infinite hypernatural (such as @{term whn})*} -lemma sumhr_hypreal_of_hypnat_omega: +lemma sumhr_hypreal_of_hypnat_omega: "sumhr(0,whn,%i. 1) = hypreal_of_hypnat whn" by (simp add: sumhr_const) @@ -108,33 +108,33 @@ (* maybe define omega = hypreal_of_hypnat whn + 1 *) apply (unfold star_class_defs omega_def hypnat_omega_def of_hypnat_def star_of_def) -apply (simp add: starfun_star_n starfun2_star_n real_of_nat_def) +apply (simp add: starfun_star_n starfun2_star_n) done -lemma sumhr_minus_one_realpow_zero [simp]: +lemma sumhr_minus_one_realpow_zero [simp]: "!!N. sumhr(0, N + N, %i. (-1) ^ (i+1)) = 0" unfolding sumhr_app -apply transfer +apply transfer apply (simp del: power_Suc add: mult_2 [symmetric]) apply (induct_tac N) apply simp_all done lemma sumhr_interval_const: - "(\n. m \ Suc n --> f n = r) & m \ na - ==> sumhr(hypnat_of_nat m,hypnat_of_nat na,f) = + "(\n. m \ Suc n --> f n = r) & m \ na + ==> sumhr(hypnat_of_nat m,hypnat_of_nat na,f) = (hypreal_of_nat (na - m) * hypreal_of_real r)" unfolding sumhr_app by transfer simp lemma starfunNat_sumr: "!!N. ( *f* (%n. setsum f {0.. abs (sumhr(M, N, f)) @= 0" apply (cut_tac x = M and y = N in linorder_less_linear) apply (auto simp add: approx_refl) apply (drule approx_sym [THEN approx_minus_iff [THEN iffD1]]) -apply (auto dest: approx_hrabs +apply (auto dest: approx_hrabs simp add: sumhr_split_diff) done @@ -166,7 +166,7 @@ by (auto simp add: sums_NSsums_iff [symmetric] not_le[symmetric] intro!: sums_finite) lemma NSsummable_NSCauchy: - "NSsummable f = + "NSsummable f = (\M \ HNatInfinite. \N \ HNatInfinite. abs (sumhr(M,N,f)) @= 0)" apply (auto simp add: summable_NSsummable_iff [symmetric] summable_iff_convergent convergent_NSconvergent_iff atLeast0LessThan[symmetric] @@ -175,7 +175,7 @@ apply auto apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym]) apply (rule_tac [2] approx_minus_iff [THEN iffD2]) -apply (auto dest: approx_hrabs_zero_cancel +apply (auto dest: approx_hrabs_zero_cancel simp add: sumhr_split_diff atLeast0LessThan[symmetric]) done diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/HyperDef.thy --- a/src/HOL/NSA/HyperDef.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/HyperDef.thy Tue Nov 10 14:18:41 2015 +0000 @@ -231,19 +231,21 @@ text{*A few lemmas first*} -lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} | - (\y. {n::nat. x = real n} = {y})" +lemma lemma_omega_empty_singleton_disj: + "{n::nat. x = real n} = {} \ (\y. {n::nat. x = real n} = {y})" by force lemma lemma_finite_omega_set: "finite {n::nat. x = real n}" -by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) + using lemma_omega_empty_singleton_disj [of x] by auto lemma not_ex_hypreal_of_real_eq_omega: "~ (\x. hypreal_of_real x = omega)" -apply (simp add: omega_def) -apply (simp add: star_of_def star_n_eq_iff) -apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] - lemma_finite_omega_set [THEN FreeUltrafilterNat.finite]) +apply (simp add: omega_def star_of_def star_n_eq_iff) +apply clarify +apply (rule_tac x2="x-1" in lemma_finite_omega_set [THEN FreeUltrafilterNat.finite, THEN notE]) +apply (rule eventually_mono) +prefer 2 apply assumption +apply auto done lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \ omega" @@ -262,7 +264,7 @@ lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\x. hypreal_of_real x = epsilon)" by (auto simp add: epsilon_def star_of_def star_n_eq_iff - lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite] simp del: real_of_nat_Suc) + lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite] simp del: of_nat_Suc) lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \ epsilon" by (insert not_ex_hypreal_of_real_eq_epsilon, auto) @@ -307,21 +309,8 @@ (* is a hyperreal c.f. NS extension *) (*------------------------------------------------------------*) -lemma hypreal_of_nat_eq: - "hypreal_of_nat (n::nat) = hypreal_of_real (real n)" -by (simp add: real_of_nat_def) - -lemma hypreal_of_nat: - "hypreal_of_nat m = star_n (%n. real m)" -apply (fold star_of_def) -apply (simp add: real_of_nat_def) -done - -(* -FIXME: we should declare this, as for type int, but many proofs would break. -It replaces x+-y by x-y. -Addsimps [symmetric hypreal_diff_def] -*) +lemma hypreal_of_nat: "hypreal_of_nat m = star_n (%n. real m)" +by (simp add: star_of_def [symmetric]) declaration {* K (Lin_Arith.add_inj_thms [@{thm star_of_le} RS iffD2, @@ -383,12 +372,6 @@ lemma two_hrealpow_ge_one [simp]: "(1::hypreal) \ 2 ^ n" by (insert power_increasing [of 0 n "2::hypreal"], simp) -lemma two_hrealpow_gt [simp]: "hypreal_of_nat n < 2 ^ n" -apply (induct n) -apply (auto simp add: distrib_right) -apply (cut_tac n = n in two_hrealpow_ge_one, arith) -done - lemma hrealpow: "star_n X ^ m = star_n (%n. (X n::real) ^ m)" apply (induct_tac "m") diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/NSA.thy --- a/src/HOL/NSA/NSA.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/NSA.thy Tue Nov 10 14:18:41 2015 +0000 @@ -2022,7 +2022,7 @@ lemma lemma_Infinitesimal: "(\r. 0 < r --> x < r) = (\n. x < inverse(real (Suc n)))" -apply (auto simp add: real_of_nat_Suc_gt_zero simp del: real_of_nat_Suc) +apply (auto simp add: real_of_nat_Suc_gt_zero simp del: of_nat_Suc) apply (blast dest!: reals_Archimedean intro: order_less_trans) done @@ -2034,10 +2034,10 @@ apply (simp (no_asm_use)) apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN star_of_less [THEN iffD2], THEN [2] impE]) prefer 2 apply assumption - apply (simp add: real_of_nat_def) -apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: real_of_nat_Suc) + apply simp +apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: of_nat_Suc) apply (drule star_of_less [THEN iffD2]) -apply (simp add: real_of_nat_def) +apply simp apply (blast intro: order_less_trans) done @@ -2057,16 +2057,11 @@ by (auto simp add: less_Suc_eq) (*------------------------------------------- - Prove that any segment is finite and - hence cannot belong to FreeUltrafilterNat + Prove that any segment is finite and hence cannot belong to FreeUltrafilterNat -------------------------------------------*) -lemma finite_nat_segment: "finite {n::nat. n < m}" -apply (induct "m") -apply (auto simp add: Suc_Un_eq) -done lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}" -by (auto intro: finite_nat_segment) + by (auto intro: finite_Collect_less_nat) lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}" apply (cut_tac x = u in reals_Archimedean2, safe) @@ -2106,13 +2101,12 @@ text{*@{term omega} is a member of @{term HInfinite}*} -lemma FreeUltrafilterNat_omega: "eventually (\n. u < real n) FreeUltrafilterNat" - by (fact FreeUltrafilterNat_nat_gt_real) - theorem HInfinite_omega [simp]: "omega \ HInfinite" apply (simp add: omega_def) apply (rule FreeUltrafilterNat_HInfinite) -apply (simp add: real_of_nat_Suc diff_less_eq [symmetric] FreeUltrafilterNat_omega) +apply clarify +apply (rule_tac u1 = "u-1" in eventually_mono [OF _ FreeUltrafilterNat_nat_gt_real]) +apply auto done (*----------------------------------------------- @@ -2145,21 +2139,22 @@ lemma finite_inverse_real_of_posnat_gt_real: "0 < u ==> finite {n. u < inverse(real(Suc n))}" -apply (simp (no_asm_simp) add: real_of_nat_less_inverse_iff del: real_of_nat_Suc) -apply (simp (no_asm_simp) add: real_of_nat_Suc less_diff_eq [symmetric]) -apply (rule finite_real_of_nat_less_real) -done +proof (simp only: real_of_nat_less_inverse_iff) + have "{n. 1 + real n < inverse u} = {n. real n < inverse u - 1}" + by fastforce + thus "finite {n. real (Suc n) < inverse u}" + using finite_real_of_nat_less_real [of "inverse u - 1"] by auto +qed lemma lemma_real_le_Un_eq2: "{n. u \ inverse(real(Suc n))} = {n. u < inverse(real(Suc n))} Un {n. u = inverse(real(Suc n))}" -apply (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) -done +by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) lemma finite_inverse_real_of_posnat_ge_real: "0 < u ==> finite {n. u \ inverse(real(Suc n))}" by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_epsilon_set finite_inverse_real_of_posnat_gt_real - simp del: real_of_nat_Suc) + simp del: of_nat_Suc) lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat: "0 < u ==> \ eventually (\n. u \ inverse(real(Suc n))) FreeUltrafilterNat" @@ -2187,7 +2182,7 @@ lemma SEQ_Infinitesimal: "( *f* (%n::nat. inverse(real(Suc n)))) whn : Infinitesimal" by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff - real_of_nat_Suc_gt_zero FreeUltrafilterNat_inverse_real_of_posnat del: real_of_nat_Suc) + real_of_nat_Suc_gt_zero FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc) text{* Example where we get a hyperreal from a real sequence for which a particular property holds. The theorem is diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/NSComplex.thy --- a/src/HOL/NSA/NSComplex.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/NSComplex.thy Tue Nov 10 14:18:41 2015 +0000 @@ -552,7 +552,6 @@ lemma NSDeMoivre: "!!a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)" apply transfer -apply (fold real_of_nat_def) apply (rule DeMoivre) done @@ -563,15 +562,15 @@ lemma NSDeMoivre_ext: "!!a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)" -by transfer (fold real_of_nat_def, rule DeMoivre) +by transfer (rule DeMoivre) lemma NSDeMoivre2: "!!a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)" -by transfer (fold real_of_nat_def, rule DeMoivre2) +by transfer (rule DeMoivre2) lemma DeMoivre2_ext: "!! a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)" -by transfer (fold real_of_nat_def, rule DeMoivre2) +by transfer (rule DeMoivre2) lemma hcis_inverse [simp]: "!!a. inverse(hcis a) = hcis (-a)" by transfer (rule cis_inverse) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/NatStar.thy --- a/src/HOL/NSA/NatStar.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/NatStar.thy Tue Nov 10 14:18:41 2015 +0000 @@ -115,7 +115,7 @@ @{term real_of_nat} *} lemma starfunNat_real_of_nat: "( *f* real) = hypreal_of_hypnat" -by transfer (simp add: fun_eq_iff real_of_nat_def) +by transfer (simp add: fun_eq_iff) lemma starfun_inverse_real_of_nat_eq: "N \ HNatInfinite diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NSA/Star.thy --- a/src/HOL/NSA/Star.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NSA/Star.thy Tue Nov 10 14:18:41 2015 +0000 @@ -110,7 +110,7 @@ (* Prove that abs for hypreal is a nonstandard extension of abs for real w/o use of congruence property (proved after this for general - nonstandard extensions of real valued functions). + nonstandard extensions of real valued functions). Proof now Uses the ultrafilter tactic! *) @@ -307,12 +307,12 @@ "(star_n X \ Infinitesimal) = (\m. eventually (\n. norm(X n) < inverse(real(Suc m))) \)" by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def hnorm_def star_of_nat_def starfun_star_n - star_n_inverse star_n_less real_of_nat_def) + star_n_inverse star_n_less) lemma HNatInfinite_inverse_Infinitesimal [simp]: "n \ HNatInfinite ==> inverse (hypreal_of_hypnat n) \ Infinitesimal" apply (cases n) -apply (auto simp add: of_hypnat_def starfun_star_n real_of_nat_def [symmetric] star_n_inverse real_norm_def +apply (auto simp add: of_hypnat_def starfun_star_n star_n_inverse real_norm_def HNatInfinite_FreeUltrafilterNat_iff Infinitesimal_FreeUltrafilterNat_iff2) apply (drule_tac x="Suc m" in spec) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/NthRoot.thy Tue Nov 10 14:18:41 2015 +0000 @@ -16,7 +16,7 @@ lemma inverse_sgn: "sgn (inverse a) = inverse (sgn a :: real)" by (simp add: sgn_real_def) -lemma power_eq_iff_eq_base: +lemma power_eq_iff_eq_base: fixes a b :: "_ :: linordered_semidom" shows "0 < n \ 0 \ a \ 0 \ b \ a ^ n = b ^ n \ a = b" using power_eq_imp_eq_base[of a n b] by auto @@ -274,7 +274,7 @@ lemma continuous_real_root[continuous_intros]: "continuous F f \ continuous F (\x. root n (f x))" unfolding continuous_def by (rule tendsto_real_root) - + lemma continuous_on_real_root[continuous_intros]: "continuous_on s f \ continuous_on s (\x. root n (f x))" unfolding continuous_on_def by (auto intro: tendsto_real_root) @@ -326,7 +326,7 @@ qed next show "DERIV (\x. - (x ^ n)) (root n x) :> - real n * root n x ^ (n - Suc 0)" - by (auto intro!: derivative_eq_intros simp: real_of_nat_def) + by (auto intro!: derivative_eq_intros) show "- real n * root n x ^ (n - Suc 0) \ 0" using n x by simp qed (rule isCont_real_root) @@ -463,7 +463,7 @@ lemma continuous_real_sqrt[continuous_intros]: "continuous F f \ continuous F (\x. sqrt (f x))" unfolding sqrt_def by (rule continuous_real_root) - + lemma continuous_on_real_sqrt[continuous_intros]: "continuous_on s f \ continuous_on s (\x. sqrt (f x))" unfolding sqrt_def by (rule continuous_on_real_root) @@ -510,14 +510,14 @@ proof cases assume "x=0" thus ?thesis by simp next - assume nz: "x\0" + assume nz: "x\0" hence pos: "0 0" by (simp add: divide_inverse nneg nz) + proof (rule right_inverse_eq [THEN iffD1, THEN sym]) + show "sqrt x / x \ 0" by (simp add: divide_inverse nneg nz) show "inverse (sqrt x) / (sqrt x / x) = 1" - by (simp add: divide_inverse mult.assoc [symmetric] - power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) + by (simp add: divide_inverse mult.assoc [symmetric] + power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) qed qed @@ -537,7 +537,7 @@ lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u" by (simp add: divide_less_eq) -lemma four_x_squared: +lemma four_x_squared: fixes x::real shows "4 * x\<^sup>2 = (2 * x)\<^sup>2" by (simp add: power2_eq_square) @@ -548,7 +548,7 @@ subsection \Square Root of Sum of Squares\ -lemma sum_squares_bound: +lemma sum_squares_bound: fixes x:: "'a::linordered_field" shows "2*x*y \ x^2 + y^2" proof - @@ -560,14 +560,14 @@ by arith qed -lemma arith_geo_mean: +lemma arith_geo_mean: fixes u:: "'a::linordered_field" assumes "u\<^sup>2 = x*y" "x\0" "y\0" shows "u \ (x + y)/2" apply (rule power2_le_imp_le) using sum_squares_bound assms apply (auto simp: zero_le_mult_iff) by (auto simp: algebra_simps power2_eq_square) -lemma arith_geo_mean_sqrt: +lemma arith_geo_mean_sqrt: fixes x::real assumes "x\0" "y\0" shows "sqrt(x*y) \ (x + y)/2" apply (rule arith_geo_mean) using assms @@ -645,10 +645,10 @@ using less_eq_real_def sqrt2_less_2 apply force apply assumption apply (rule le_less_trans [where y = "y*2"]) - using less_eq_real_def sqrt2_less_2 mult_le_cancel_left - apply auto + using less_eq_real_def sqrt2_less_2 mult_le_cancel_left + apply auto done - + lemma LIMSEQ_root: "(\n. root n n) ----> 1" proof - def x \ "\n. root n n - 1" @@ -660,7 +660,7 @@ { fix n :: nat assume "2 < n" have "1 + (real (n - 1) * n) / 2 * x n^2 = 1 + of_nat (n choose 2) * x n^2" using \2 < n\ unfolding gbinomial_def binomial_gbinomial - by (simp add: atLeast0AtMost atMost_Suc field_simps real_of_nat_diff numeral_2_eq_2 real_eq_of_nat[symmetric]) + by (simp add: atLeast0AtMost atMost_Suc field_simps of_nat_diff numeral_2_eq_2) also have "\ \ (\k\{0, 2}. of_nat (n choose k) * x n^k)" by (simp add: x_def) also have "\ \ (\k=0..n. of_nat (n choose k) * x n^k)" @@ -697,7 +697,7 @@ (simp_all add: at_infinity_eq_at_top_bot) { fix n :: nat assume "1 < n" have "1 + x n * n = 1 + of_nat (n choose 1) * x n^1" - using \1 < n\ unfolding gbinomial_def binomial_gbinomial by (simp add: real_eq_of_nat[symmetric]) + using \1 < n\ unfolding gbinomial_def binomial_gbinomial by simp also have "\ \ (\k\{0, 1}. of_nat (n choose k) * x n^k)" by (simp add: x_def) also have "\ \ (\k=0..n. of_nat (n choose k) * x n^k)" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Old_Number_Theory/Fib.thy --- a/src/HOL/Old_Number_Theory/Fib.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Old_Number_Theory/Fib.thy Tue Nov 10 14:18:41 2015 +0000 @@ -86,9 +86,7 @@ (if n mod 2 = 0 then fib (Suc n) * fib (Suc n) - 1 else fib (Suc n) * fib (Suc n) + 1)" apply (rule int_int_eq [THEN iffD1]) - apply (simp add: fib_Cassini_int) - apply (subst of_nat_diff) - apply (insert fib_Suc_gr_0 [of n], simp_all) + using fib_Cassini_int apply (auto simp add: Suc_leI fib_Suc_gr_0 of_nat_diff) done diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy --- a/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy Tue Nov 10 14:18:41 2015 +0000 @@ -274,7 +274,7 @@ lemma S_card: "((p - 1) div 2) * ((q - 1) div 2) = int (card(S))" using P_set_card Q_set_card P_set_finite Q_set_finite - by (auto simp add: S_def zmult_int) + by (simp add: S_def) lemma S1_Int_S2_prop: "S1 \ S2 = {}" by (auto simp add: S1_def S2_def) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Tue Nov 10 14:18:41 2015 +0000 @@ -717,7 +717,7 @@ apply (simp add: fin_Pair emeasure_count_space X_subset fin_X) apply (subst nn_integral_count_space) using A apply simp - apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric]) + apply (simp del: of_nat_setsum add: of_nat_setsum[symmetric]) apply (subst card_gt_0_iff) apply (simp add: fin_Pair) apply (subst card_SigmaI[symmetric]) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Bochner_Integration.thy Tue Nov 10 14:18:41 2015 +0000 @@ -153,7 +153,7 @@ qed qed -lemma real_indicator: "real (indicator A x :: ereal) = indicator A x" +lemma real_indicator: "real_of_ereal (indicator A x :: ereal) = indicator A x" unfolding indicator_def by auto lemma split_indicator_asm: @@ -182,9 +182,9 @@ sup: "\x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\i x. 0 \ U i x" by blast - def U' \ "\i x. indicator (space M) x * real (U i x)" + def U' \ "\i x. indicator (space M) x * real_of_ereal (U i x)" then have U'_sf[measurable]: "\i. simple_function M (U' i)" - using U by (auto intro!: simple_function_compose1[where g=real]) + using U by (auto intro!: simple_function_compose1[where g=real_of_ereal]) show "P u" proof (rule seq) @@ -209,7 +209,7 @@ by simp next fix i - have "U' i ` space M \ real ` (U i ` space M)" "finite (U i ` space M)" + have "U' i ` space M \ real_of_ereal ` (U i ` space M)" "finite (U i ` space M)" unfolding U'_def using U(1) by (auto dest: simple_functionD) then have fin: "finite (U' i ` space M)" by (metis finite_subset finite_imageI) @@ -472,7 +472,7 @@ emeasure M ((\x. ereal (f x)) -` {ereal (f y)} \ space M)" by simp } with f have "simple_bochner_integral M f = (\\<^sup>Sx. f x \M)" unfolding simple_integral_def - by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ereal (f x)" and v=real]) + by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ereal (f x)" and v=real_of_ereal]) (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases intro!: setsum.cong ereal_cong_mult simp: setsum_ereal[symmetric] times_ereal.simps(1)[symmetric] ac_simps @@ -1644,7 +1644,7 @@ lemma integral_eq_nn_integral: assumes [measurable]: "f \ borel_measurable M" assumes nonneg: "AE x in M. 0 \ f x" - shows "integral\<^sup>L M f = real (\\<^sup>+ x. ereal (f x) \M)" + shows "integral\<^sup>L M f = real_of_ereal (\\<^sup>+ x. ereal (f x) \M)" proof cases assume *: "(\\<^sup>+ x. ereal (f x) \M) = \" also have "(\\<^sup>+ x. ereal (f x) \M) = (\\<^sup>+ x. ereal (norm (f x)) \M)" @@ -2138,15 +2138,15 @@ lemma real_lebesgue_integral_def: assumes f[measurable]: "integrable M f" - shows "integral\<^sup>L M f = real (\\<^sup>+x. f x \M) - real (\\<^sup>+x. - f x \M)" + shows "integral\<^sup>L M f = real_of_ereal (\\<^sup>+x. f x \M) - real_of_ereal (\\<^sup>+x. - f x \M)" proof - have "integral\<^sup>L M f = integral\<^sup>L M (\x. max 0 (f x) - max 0 (- f x))" by (auto intro!: arg_cong[where f="integral\<^sup>L M"]) also have "\ = integral\<^sup>L M (\x. max 0 (f x)) - integral\<^sup>L M (\x. max 0 (- f x))" by (intro integral_diff integrable_max integrable_minus integrable_zero f) - also have "integral\<^sup>L M (\x. max 0 (f x)) = real (\\<^sup>+x. max 0 (f x) \M)" + also have "integral\<^sup>L M (\x. max 0 (f x)) = real_of_ereal (\\<^sup>+x. max 0 (f x) \M)" by (subst integral_eq_nn_integral[symmetric]) auto - also have "integral\<^sup>L M (\x. max 0 (- f x)) = real (\\<^sup>+x. max 0 (- f x) \M)" + also have "integral\<^sup>L M (\x. max 0 (- f x)) = real_of_ereal (\\<^sup>+x. max 0 (- f x) \M)" by (subst integral_eq_nn_integral[symmetric]) auto also have "(\x. ereal (max 0 (f x))) = (\x. max 0 (ereal (f x)))" by (auto simp: max_def) @@ -2346,13 +2346,13 @@ lemma (in finite_measure) ereal_integral_real: assumes [measurable]: "f \ borel_measurable M" assumes ae: "AE x in M. 0 \ f x" "AE x in M. f x \ ereal B" - shows "ereal (\x. real (f x) \M) = (\\<^sup>+x. f x \M)" + shows "ereal (\x. real_of_ereal (f x) \M) = (\\<^sup>+x. f x \M)" proof (subst nn_integral_eq_integral[symmetric]) - show "integrable M (\x. real (f x))" + show "integrable M (\x. real_of_ereal (f x))" using ae by (intro integrable_const_bound[where B=B]) (auto simp: real_le_ereal_iff) - show "AE x in M. 0 \ real (f x)" + show "AE x in M. 0 \ real_of_ereal (f x)" using ae by (auto simp: real_of_ereal_pos) - show "(\\<^sup>+ x. ereal (real (f x)) \M) = integral\<^sup>N M f" + show "(\\<^sup>+ x. ereal (real_of_ereal (f x)) \M) = integral\<^sup>N M f" using ae by (intro nn_integral_cong_AE) (auto simp: ereal_real) qed diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Tue Nov 10 14:18:41 2015 +0000 @@ -618,7 +618,7 @@ interpret sigma_algebra UNIV ?SIGMA by (intro sigma_algebra_sigma_sets) simp_all have *: "?set = (\n. UNIV - {x::'a. x \ i < a + 1 / real (Suc n)})" - proof (safe, simp_all add: not_less del: real_of_nat_Suc) + proof (safe, simp_all add: not_less del: of_nat_Suc) fix x :: 'a assume "a < x \ i" with reals_Archimedean[of "x \ i - a"] obtain n where "a + 1 / real (Suc n) < x \ i" @@ -655,7 +655,7 @@ fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" then have i: "i \ Basis" by auto have *: "{x::'a. x\i < a} = (\n. {x. x\i \ a - 1/real (Suc n)})" - proof (safe, simp_all del: real_of_nat_Suc) + proof (safe, simp_all del: of_nat_Suc) fix x::'a assume *: "x\i < a" with reals_Archimedean[of "a - x\i"] obtain n where "x \ i < a - 1 / (real (Suc n))" @@ -1055,7 +1055,7 @@ lemma measurable_real_floor[measurable]: "(floor :: real \ int) \ measurable borel (count_space UNIV)" proof - - have "\a x. \x\ = a \ (real a \ x \ x < real (a + 1))" + have "\a x. \x\ = a \ (real_of_int a \ x \ x < real_of_int (a + 1))" by (auto intro: floor_eq2) then show ?thesis by (auto simp: vimage_def measurable_count_space_eq2_countable) @@ -1065,7 +1065,7 @@ "(ceiling :: real \ int) \ measurable borel (count_space UNIV)" unfolding ceiling_def[abs_def] by simp -lemma borel_measurable_real_floor: "(\x::real. real \x\) \ borel_measurable borel" +lemma borel_measurable_real_floor: "(\x::real. real_of_int \x\) \ borel_measurable borel" by simp lemma borel_measurable_root [measurable]: "root n \ borel_measurable borel" @@ -1116,7 +1116,7 @@ lemma borel_measurable_real_of_ereal[measurable (raw)]: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" - shows "(\x. real (f x)) \ borel_measurable M" + shows "(\x. real_of_ereal (f x)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{\, -\ }"]) apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) @@ -1125,10 +1125,10 @@ lemma borel_measurable_ereal_cases: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" - assumes H: "(\x. H (ereal (real (f x)))) \ borel_measurable M" + assumes H: "(\x. H (ereal (real_of_ereal (f x)))) \ borel_measurable M" shows "(\x. H (f x)) \ borel_measurable M" proof - - let ?F = "\x. if f x = \ then H \ else if f x = - \ then H (-\) else H (ereal (real (f x)))" + let ?F = "\x. if f x = \ then H \ else if f x = - \ then H (-\) else H (ereal (real_of_ereal (f x)))" { fix x have "H (f x) = ?F x" by (cases "f x") auto } with f H show ?thesis by simp qed @@ -1150,15 +1150,15 @@ fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" - assumes H: "{x \ space M. H (ereal (real (f x))) (ereal (real (g x)))} \ sets M" + assumes H: "{x \ space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \ sets M" "{x \ space borel. H (-\) (ereal x)} \ sets borel" "{x \ space borel. H (\) (ereal x)} \ sets borel" "{x \ space borel. H (ereal x) (-\)} \ sets borel" "{x \ space borel. H (ereal x) (\)} \ sets borel" shows "{x \ space M. H (f x) (g x)} \ sets M" proof - - let ?G = "\y x. if g x = \ then H y \ else if g x = -\ then H y (-\) else H y (ereal (real (g x)))" - let ?F = "\x. if f x = \ then ?G \ x else if f x = -\ then ?G (-\) x else ?G (ereal (real (f x))) x" + let ?G = "\y x. if g x = \ then H y \ else if g x = -\ then H y (-\) else H y (ereal (real_of_ereal (g x)))" + let ?F = "\x. if f x = \ then ?G \ x else if f x = -\ then ?G (-\) x else ?G (ereal (real_of_ereal (f x))) x" { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } note * = this from assms show ?thesis @@ -1180,12 +1180,12 @@ lemma borel_measurable_ereal_iff_real: fixes f :: "'a \ ereal" shows "f \ borel_measurable M \ - ((\x. real (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" + ((\x. real_of_ereal (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" proof safe - assume *: "(\x. real (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" "f -` {-\} \ space M \ sets M" + assume *: "(\x. real_of_ereal (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" "f -` {-\} \ space M \ sets M" have "f -` {\} \ space M = {x\space M. f x = \}" "f -` {-\} \ space M = {x\space M. f x = -\}" by auto with * have **: "{x\space M. f x = \} \ sets M" "{x\space M. f x = -\} \ sets M" by simp_all - let ?f = "\x. if f x = \ then \ else if f x = -\ then -\ else ereal (real (f x))" + let ?f = "\x. if f x = \ then \ else if f x = -\ then -\ else ereal (real_of_ereal (f x))" have "?f \ borel_measurable M" using * ** by (intro measurable_If) auto also have "?f = f" by (auto simp: fun_eq_iff ereal_real) finally show "f \ borel_measurable M" . @@ -1221,15 +1221,15 @@ fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" - assumes H: "(\x. H (ereal (real (f x))) (ereal (real (g x)))) \ borel_measurable M" - "(\x. H (-\) (ereal (real (g x)))) \ borel_measurable M" - "(\x. H (\) (ereal (real (g x)))) \ borel_measurable M" - "(\x. H (ereal (real (f x))) (-\)) \ borel_measurable M" - "(\x. H (ereal (real (f x))) (\)) \ borel_measurable M" + assumes H: "(\x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \ borel_measurable M" + "(\x. H (-\) (ereal (real_of_ereal (g x)))) \ borel_measurable M" + "(\x. H (\) (ereal (real_of_ereal (g x)))) \ borel_measurable M" + "(\x. H (ereal (real_of_ereal (f x))) (-\)) \ borel_measurable M" + "(\x. H (ereal (real_of_ereal (f x))) (\)) \ borel_measurable M" shows "(\x. H (f x) (g x)) \ borel_measurable M" proof - - let ?G = "\y x. if g x = \ then H y \ else if g x = - \ then H y (-\) else H y (ereal (real (g x)))" - let ?F = "\x. if f x = \ then ?G \ x else if f x = - \ then ?G (-\) x else ?G (ereal (real (f x))) x" + let ?G = "\y x. if g x = \ then H y \ else if g x = - \ then H y (-\) else H y (ereal (real_of_ereal (g x)))" + let ?F = "\x. if f x = \ then ?G \ x else if f x = - \ then ?G (-\) x else ?G (ereal (real_of_ereal (f x))) x" { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } note * = this from assms show ?thesis unfolding * by simp @@ -1427,7 +1427,7 @@ also have "\ < ?I i / 2 + ?I i / 2" by (intro add_strict_mono d less_trans[OF _ j] *) also have "\ \ ?I i" - by (simp add: field_simps real_of_nat_Suc) + by (simp add: field_simps of_nat_Suc) finally show "dist (f y) (f z) \ ?I i" by simp qed diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Distributions.thy --- a/src/HOL/Probability/Distributions.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Distributions.thy Tue Nov 10 14:18:41 2015 +0000 @@ -25,7 +25,7 @@ by simp show "random_variable lborel (\x. t + c * X x)" by simp - + have "AE x in lborel. 0 \ f x" using f by (simp add: distributed_def) from AE_borel_affine[OF _ _ this, where c="1/c" and t="- t / c"] c @@ -34,7 +34,7 @@ have eq: "\x. ereal \c\ * (f x / ereal \c\) = f x" using c by (simp add: divide_ereal_def ac_simps one_ereal_def[symmetric]) - + have "density lborel f = distr M lborel X" using f by (simp add: distributed_def) with c show "distr M lborel (\x. t + c * X x) = density lborel (\x. f ((x - t) / c) / ereal \c\)" @@ -78,7 +78,7 @@ proof - let ?f = "\k x. x^k * exp (-x) / fact k" let ?F = "\k x. - (\n\k. (x^n * exp (-x)) / fact n)" - have "?I * (inverse (real_of_nat (fact k))) = + have "?I * (inverse (real_of_nat (fact k))) = (\\<^sup>+x. ereal (x^k * exp (-x)) * indicator {0 .. a} x * (inverse (real_of_nat (fact k))) \lborel)" by (intro nn_integral_multc[symmetric]) auto also have "\ = (\\<^sup>+x. ereal (?f k x) * indicator {0 .. a} x \lborel)" @@ -97,12 +97,12 @@ ?f k x - ((real (Suc k) - x) * x ^ k * exp (- x)) / (fact (Suc k))" by (intro DERIV_diff Suc) (auto intro!: derivative_eq_intros simp del: fact_Suc power_Suc - simp add: field_simps power_Suc[symmetric] real_of_nat_def[symmetric]) + simp add: field_simps power_Suc[symmetric]) also have "(\x. ?F k x - (x^Suc k * exp (-x)) / fact (Suc k)) = ?F (Suc k)" by simp also have "?f k x - ((real (Suc k) - x) * x ^ k * exp (- x)) / (fact (Suc k)) = ?f (Suc k) x" by (auto simp: field_simps simp del: fact_Suc) - (simp_all add: real_of_nat_Suc field_simps) + (simp_all add: of_nat_Suc field_simps) finally show ?case . qed qed auto @@ -166,20 +166,20 @@ by (auto simp: erlang_CDF_def) finally show ?thesis . next - assume "\ 0 \ a" + assume "\ 0 \ a" moreover then have "(\\<^sup>+ x. ereal (erlang_density k l x) * indicator {.. a} x \lborel) = (\\<^sup>+x. 0 \(lborel::real measure))" by (intro nn_integral_cong) (auto simp: erlang_density_def) ultimately show ?thesis by (simp add: erlang_CDF_def) qed -lemma emeasure_erlang_density: +lemma emeasure_erlang_density: "0 < l \ emeasure (density lborel (erlang_density k l)) {.. a} = erlang_CDF k l a" by (simp add: emeasure_density nn_integral_erlang_density) -lemma nn_integral_erlang_ith_moment: +lemma nn_integral_erlang_ith_moment: fixes k i :: nat and l :: real - assumes [arith]: "0 < l" + assumes [arith]: "0 < l" shows "(\\<^sup>+ x. ereal (erlang_density k l x * x ^ i) \lborel) = fact (k + i) / (fact k * l ^ i)" proof - have eq: "\x. indicator {0..} (x / l) = indicator {0..} x" @@ -241,7 +241,7 @@ shows "distributed M lborel X (erlang_density k l)" proof (rule distributedI_borel_atMost) fix a :: real - { assume "a \ 0" + { assume "a \ 0" with X have "emeasure M {x\space M. X x \ a} \ emeasure M {x\space M. X x \ 0}" by (intro emeasure_mono) auto also have "... = 0" by (auto intro!: erlang_CDF_at0 simp: X_distr[of 0]) @@ -266,31 +266,31 @@ distributed_measurable[of M lborel X "erlang_density k l"] emeasure_erlang_density[of l] erlang_distributed_le[of X k l] - by (auto intro!: erlang_distributedI simp: one_ereal_def emeasure_eq_measure) + by (auto intro!: erlang_distributedI simp: one_ereal_def emeasure_eq_measure) lemma (in prob_space) erlang_distributed_mult_const: assumes erlX: "distributed M lborel X (erlang_density k l)" assumes a_pos[arith]: "0 < \" "0 < l" shows "distributed M lborel (\x. \ * X x) (erlang_density k (l / \))" proof (subst erlang_distributed_iff, safe) - have [measurable]: "random_variable borel X" and [arith]: "0 < l " + have [measurable]: "random_variable borel X" and [arith]: "0 < l " and [simp]: "\a. 0 \ a \ prob {x \ space M. X x \ a} = erlang_CDF k l a" by(insert erlX, auto simp: erlang_distributed_iff) - show "random_variable borel (\x. \ * X x)" "0 < l / \" "0 < l / \" + show "random_variable borel (\x. \ * X x)" "0 < l / \" "0 < l / \" by (auto simp:field_simps) - + fix a:: real assume [arith]: "0 \ a" - obtain b:: real where [simp, arith]: "b = a/ \" by blast + obtain b:: real where [simp, arith]: "b = a/ \" by blast have [arith]: "0 \ b" by (auto simp: divide_nonneg_pos) - + have "prob {x \ space M. \ * X x \ a} = prob {x \ space M. X x \ b}" by (rule arg_cong[where f= prob]) (auto simp:field_simps) - + moreover have "prob {x \ space M. X x \ b} = erlang_CDF k l b" by auto moreover have "erlang_CDF k (l / \) a = erlang_CDF k l b" unfolding erlang_CDF_def by auto - ultimately show "prob {x \ space M. \ * X x \ a} = erlang_CDF k (l / \) a" by fastforce + ultimately show "prob {x \ space M. \ * X x \ a} = erlang_CDF k (l / \) a" by fastforce qed lemma (in prob_space) has_bochner_integral_erlang_ith_moment: @@ -324,7 +324,7 @@ show "expectation (\x. (X x)\<^sup>2) - (expectation X)\<^sup>2 = real (k + 1) / l\<^sup>2" using erlang_ith_moment[OF assms, of 1] erlang_ith_moment[OF assms, of 2] - by simp (auto simp: power2_eq_square field_simps real_of_nat_Suc) + by simp (auto simp: power2_eq_square field_simps of_nat_Suc) qed subsection {* Exponential distribution *} @@ -352,7 +352,7 @@ have "AE x in lborel. 0 \ exponential_density l x" using assms by (auto simp: distributed_real_AE) then have "AE x in lborel. x \ (0::real)" - apply eventually_elim + apply eventually_elim using `l < 0` apply (auto simp: exponential_density_def zero_le_mult_iff split: split_if_asm) done @@ -431,14 +431,14 @@ have randX: "random_variable borel X" using expX by (simp add: exponential_distributed_iff) moreover have randY: "random_variable borel Y" using expY by (simp add: exponential_distributed_iff) ultimately show "random_variable borel (\x. min (X x) (Y x))" by auto - + have "0 < l" by (rule exponential_distributed_params) fact moreover have "0 < u" by (rule exponential_distributed_params) fact ultimately show " 0 < l + u" by force fix a::real assume a[arith]: "0 \ a" - have gt1[simp]: "\

(x in M. a < X x ) = exp (- a * l)" by (rule exponential_distributedD_gt[OF expX a]) - have gt2[simp]: "\

(x in M. a < Y x ) = exp (- a * u)" by (rule exponential_distributedD_gt[OF expY a]) + have gt1[simp]: "\

(x in M. a < X x ) = exp (- a * l)" by (rule exponential_distributedD_gt[OF expX a]) + have gt2[simp]: "\

(x in M. a < Y x ) = exp (- a * u)" by (rule exponential_distributedD_gt[OF expY a]) have "\

(x in M. a < (min (X x) (Y x)) ) = \

(x in M. a < (X x) \ a < (Y x))" by (auto intro!:arg_cong[where f=prob]) @@ -450,16 +450,16 @@ have "{x \ space M. (min (X x) (Y x)) \a } = (space M - {x \ space M. a<(min (X x) (Y x)) })" by auto then have "1 - prob {x \ space M. a < (min (X x) (Y x))} = prob {x \ space M. (min (X x) (Y x)) \ a}" - using randX randY by (auto simp: prob_compl) + using randX randY by (auto simp: prob_compl) then show "prob {x \ space M. (min (X x) (Y x)) \ a} = 1 - exp (- a * (l + u))" using indep_prob by auto qed - + lemma (in prob_space) exponential_distributed_Min: assumes finI: "finite I" assumes A: "I \ {}" assumes expX: "\i. i \ I \ distributed M lborel (X i) (exponential_density (l i))" - assumes ind: "indep_vars (\i. borel) X I" + assumes ind: "indep_vars (\i. borel) X I" shows "distributed M lborel (\x. Min ((\i. X i x)`I)) (exponential_density (\i\I. l i))" using assms proof (induct rule: finite_ne_induct) @@ -468,7 +468,7 @@ case (insert i I) then have "distributed M lborel (\x. min (X i x) (Min ((\i. X i x)`I))) (exponential_density (l i + (\i\I. l i)))" by (intro exponential_distributed_min indep_vars_Min insert) - (auto intro: indep_vars_subset) + (auto intro: indep_vars_subset) then show ?case using insert by simp qed @@ -499,22 +499,22 @@ done next note zero_le_mult_iff[simp] zero_le_divide_iff[simp] - + have I_eq1: "integral\<^sup>N lborel (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l) = 1" using nn_integral_erlang_ith_moment[of l "Suc k\<^sub>1 + Suc k\<^sub>2 - 1" 0] by (simp del: fact_Suc) - + have 1: "(\\<^sup>+ x. ereal (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l x * indicator {0<..} x) \lborel) = 1" apply (subst I_eq1[symmetric]) unfolding erlang_density_def by (auto intro!: nn_integral_cong split:split_indicator) - + have "prob_space (density lborel ?LHS)" unfolding times_ereal.simps[symmetric] - by (intro prob_space_convolution_density) + by (intro prob_space_convolution_density) (auto intro!: prob_space_erlang_density erlang_density_nonneg) then have 2: "integral\<^sup>N lborel ?LHS = 1" by (auto dest!: prob_space.emeasure_space_1 simp: emeasure_density) - + let ?I = "(integral\<^sup>N lborel (\y. ereal ((1 - y)^ k\<^sub>1 * y^k\<^sub>2 * indicator {0..1} y)))" let ?C = "(fact (Suc (k\<^sub>1 + k\<^sub>2))) / ((fact k\<^sub>1) * (fact k\<^sub>2))" let ?s = "Suc k\<^sub>1 + Suc k\<^sub>2 - 1" @@ -523,7 +523,7 @@ { fix x :: real assume [arith]: "0 < x" have *: "\x y n. (x - y * x::real)^n = x^n * (1 - y)^n" unfolding power_mult_distrib[symmetric] by (simp add: field_simps) - + have "?LHS x = ?L x" unfolding erlang_density_def by (auto intro!: nn_integral_cong split:split_indicator) @@ -534,7 +534,7 @@ apply (auto simp add: erlang_density_def mult_less_0_iff exp_minus field_simps exp_diff power_add * simp del: fact_Suc split: split_indicator) done - finally have "(\\<^sup>+y. ereal (erlang_density k\<^sub>1 l (x - y) * erlang_density k\<^sub>2 l y) \lborel) = + finally have "(\\<^sup>+y. ereal (erlang_density k\<^sub>1 l (x - y) * erlang_density k\<^sub>2 l y) \lborel) = (\x. ereal ?C * ?I * erlang_density ?s l x) x" by simp } note * = this @@ -548,9 +548,9 @@ by (auto intro!: nn_integral_cong simp: * split: split_indicator) also have "... = ereal (?C) * ?I" using 1 - by (auto simp: nn_integral_nonneg nn_integral_cmult) + by (auto simp: nn_integral_nonneg nn_integral_cmult) finally have " ereal (?C) * ?I = 1" by presburger - + then show ?thesis using * by simp qed @@ -585,7 +585,7 @@ also have "(\x. (X i x) + (\i\ I. X i x)) = (\x. \i\insert i I. X i x)" using insert by auto also have "Suc(k i) + Suc ((\i\I. Suc (k i)) - 1) - 1 = (\i\insert i I. Suc (k i)) - 1" - using insert by (auto intro!: Suc_pred simp: ac_simps) + using insert by (auto intro!: Suc_pred simp: ac_simps) finally show ?case by fast qed @@ -593,7 +593,7 @@ assumes finI: "finite I" assumes A: "I \ {}" assumes expX: "\i. i \ I \ distributed M lborel (X i) (exponential_density l)" - assumes ind: "indep_vars (\i. borel) X I" + assumes ind: "indep_vars (\i. borel) X I" shows "distributed M lborel (\x. \i\I. X i x) (erlang_density ((card I) - 1) l)" proof - obtain i where "i \ I" using assms by auto @@ -606,22 +606,22 @@ shows "entropy b lborel X = log b (exp 1 / l)" proof - have l[simp, arith]: "0 < l" by (rule exponential_distributed_params[OF D]) - + have [simp]: "integrable lborel (exponential_density l)" using distributed_integrable[OF D, of "\_. 1"] by simp have [simp]: "integral\<^sup>L lborel (exponential_density l) = 1" using distributed_integral[OF D, of "\_. 1"] by (simp add: prob_space) - + have [simp]: "integrable lborel (\x. exponential_density l x * x)" using erlang_ith_moment_integrable[OF l D, of 1] distributed_integrable[OF D, of "\x. x"] by simp have [simp]: "integral\<^sup>L lborel (\x. exponential_density l x * x) = 1 / l" using erlang_ith_moment[OF l D, of 1] distributed_integral[OF D, of "\x. x"] by simp - + have "entropy b lborel X = - (\ x. exponential_density l x * log b (exponential_density l x) \lborel)" using D by (rule entropy_distr) - also have "(\ x. exponential_density l x * log b (exponential_density l x) \lborel) = + also have "(\ x. exponential_density l x * log b (exponential_density l x) \lborel) = (\ x. (ln l * exponential_density l x - l * (exponential_density l x * x)) / ln b \lborel)" by (intro integral_cong) (auto simp: log_def ln_mult exponential_density_def field_simps) also have "\ = (ln l - 1) / ln b" @@ -674,7 +674,7 @@ using A by simp then show "emeasure M {x\space M. X x \ a} = ereal (measure lborel (A \ {..a}) / r)" using distr by simp - + have "(\\<^sup>+ x. ereal (indicator A x / measure lborel A * indicator {..a} x) \lborel) = (\\<^sup>+ x. ereal (1 / measure lborel A) * indicator (A \ {..a}) x \lborel)" by (auto intro!: nn_integral_cong split: split_indicator) @@ -698,7 +698,7 @@ by auto then show "emeasure M {x\space M. X x \ t} = emeasure lborel ({a .. b} \ {..t}) / (b - a)" proof (elim disjE conjE) - assume "t < a" + assume "t < a" then have "emeasure M {x\space M. X x \ t} \ emeasure M {x\space M. X x \ a}" using X by (auto intro!: emeasure_mono measurable_sets) also have "\ = 0" @@ -713,7 +713,7 @@ then show ?thesis using `a \ t` `a < b` using distr[OF bnds] by (simp add: emeasure_eq_measure) next - assume "b < t" + assume "b < t" have "1 = emeasure M {x\space M. X x \ b}" using distr[of b] `a < b` by (simp add: one_ereal_def emeasure_eq_measure) also have "\ \ emeasure M {x\space M. X x \ t}" @@ -752,13 +752,13 @@ proof (rule ccontr) assume "\ a < b" then have "{a .. b} = {} \ {a .. b} = {a .. a}" by simp - with uniform_distributed_params[OF D] show False + with uniform_distributed_params[OF D] show False by (auto simp: measure_def) qed lemma (in prob_space) uniform_distributed_iff: fixes a b :: real - shows "distributed M lborel X (\x. indicator {a..b} x / measure lborel {a..b}) \ + shows "distributed M lborel X (\x. indicator {a..b} x / measure lborel {a..b}) \ (X \ borel_measurable M \ a < b \ (\t\{a .. b}. \

(x in M. X x \ t)= (t - a) / (b - a)))" using uniform_distributed_bounds[of X a b] @@ -774,7 +774,7 @@ have "a < b" using uniform_distributed_bounds[OF D] . - have "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = + have "(\ x. indicator {a .. b} x / measure lborel {a .. b} * x \lborel) = (\ x. (x / measure lborel {a .. b}) * indicator {a .. b} x \lborel)" by (intro integral_cong) auto also have "(\ x. (x / measure lborel {a .. b}) * indicator {a .. b} x \lborel) = (a + b) / 2" @@ -867,9 +867,9 @@ proof (subst nn_integral_FTC_atLeast) have "((\a. - (exp (- (a\<^sup>2 * (1 + s\<^sup>2))) / (2 + 2 * s\<^sup>2))) ---> (- (0 / (2 + 2 * s\<^sup>2)))) at_top" apply (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top]) - apply (subst mult.commute) + apply (subst mult.commute) apply (auto intro!: filterlim_tendsto_pos_mult_at_top - filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] + filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident] simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff) done then show "((\a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) ---> 0) at_top" @@ -884,7 +884,7 @@ finally have "?pI ?gauss^2 = pi / 4" by (simp add: power2_eq_square) then have "?pI ?gauss = sqrt (pi / 4)" - using power_eq_iff_eq_base[of 2 "real (?pI ?gauss)" "sqrt (pi / 4)"] + using power_eq_iff_eq_base[of 2 "real_of_ereal (?pI ?gauss)" "sqrt (pi / 4)"] nn_integral_nonneg[of lborel "\x. ?gauss x * indicator {0..} x"] by (cases "?pI ?gauss") auto also have "?pI ?gauss = (\\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \lborel)" @@ -896,8 +896,8 @@ qed lemma gaussian_moment_1: - "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x)) (1 / 2)" -proof - + "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x)) (1 / 2)" +proof - have "(\\<^sup>+x. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x) \lborel) = (\\<^sup>+x. ereal (x * exp (- x\<^sup>2)) * indicator {0..} x \lborel)" by (intro nn_integral_cong) @@ -922,10 +922,10 @@ fixes k :: nat shows gaussian_moment_even_pos: "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k))) - ((sqrt pi / 2) * (fact (2 * k) / (2 ^ (2 * k) * fact k)))" + ((sqrt pi / 2) * (fact (2 * k) / (2 ^ (2 * k) * fact k)))" (is "?even") and gaussian_moment_odd_pos: - "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k + 1))) (fact k / 2)" + "has_bochner_integral lborel (\x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k + 1))) (fact k / 2)" (is "?odd") proof - let ?M = "\k x. exp (- x\<^sup>2) * x^k :: real" @@ -967,7 +967,7 @@ (\x. indicator {0..b} x *\<^sub>R (- 2 * x * exp (- x\<^sup>2) * x ^ (Suc k)) \lborel)" by (rule integral_by_parts') (auto intro!: derivative_eq_intros b - simp: real_of_nat_def[symmetric] diff_Suc real_of_nat_Suc field_simps split: nat.split) + simp: diff_Suc of_nat_Suc field_simps split: nat.split) also have "(\x. indicator {0..b} x *\<^sub>R (- 2 * x * exp (- x\<^sup>2) * x ^ (Suc k)) \lborel) = (\x. indicator {0..b} x *\<^sub>R (- 2 * (exp (- x\<^sup>2) * x ^ (k + 2))) \lborel)" by (intro integral_cong) auto @@ -979,7 +979,7 @@ qed finally have int_M_at_top: "((?f ---> (k + 1) / 2 * (\x. indicator {0..} x *\<^sub>R ?M k x \lborel)) at_top)" by simp - + have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R ?M (k + 2) x) ((k + 1) / 2 * I)" proof (rule has_bochner_integral_monotone_convergence_at_top) fix y :: real @@ -1000,7 +1000,7 @@ also have "(real (2 * k + 1) / 2 * (sqrt pi / 2 * ((fact (2 * k)) / ((2::real)^(2*k) * fact k)))) = sqrt pi / 2 * ((fact (2 * Suc k)) / ((2::real)^(2 * Suc k) * fact (Suc k)))" apply (simp add: field_simps del: fact_Suc) - apply (simp add: real_of_nat_def of_nat_mult field_simps) + apply (simp add: of_nat_mult field_simps) done finally show ?case by simp @@ -1011,7 +1011,7 @@ case (Suc k) note step[OF this] also have "(real (2 * k + 1 + 1) / (2::real) * ((fact k) / 2)) = (fact (Suc k)) / 2" - by (simp add: field_simps real_of_nat_Suc divide_simps del: fact_Suc) (simp add: field_simps) + by (simp add: field_simps of_nat_Suc divide_simps del: fact_Suc) (simp add: field_simps) finally show ?case by simp qed (insert gaussian_moment_1, simp) @@ -1037,7 +1037,7 @@ (\x. exp (- ((sqrt 2 * \) * x)\<^sup>2 / (2*\\<^sup>2)) * ((sqrt 2 * \) * x) ^ (2 * k) / sqrt (2 * pi * \\<^sup>2))" by (auto simp: fun_eq_iff field_simps real_sqrt_power[symmetric] real_sqrt_mult real_sqrt_divide power_mult eq) - also have "((sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k))) * ((2*\\<^sup>2)^k / sqrt (2 * pi * \\<^sup>2))) = + also have "((sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k))) * ((2*\\<^sup>2)^k / sqrt (2 * pi * \\<^sup>2))) = (inverse (sqrt 2 * \) * ((fact (2 * k))) / ((2/\\<^sup>2) ^ k * (fact k)))" by (auto simp: fun_eq_iff power_mult field_simps real_sqrt_power[symmetric] real_sqrt_mult power2_eq_square) @@ -1060,7 +1060,7 @@ also have "(\x. (exp (-x\<^sup>2)*\x\^(2 * k + 1)) * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2))) = (\x. exp (- (((sqrt 2 * \) * x)\<^sup>2 / (2 * \\<^sup>2))) * \sqrt 2 * \ * x\ ^ (2 * k + 1) / sqrt (2 * pi * \\<^sup>2))" by (simp add: field_simps abs_mult real_sqrt_power[symmetric] power_mult real_sqrt_mult) - also have "(fact k * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2))) = + also have "(fact k * (2^k * \^(2 * k + 1) / sqrt (pi * \\<^sup>2))) = (inverse (sqrt 2) * inverse \ * (2 ^ k * (\ * \ ^ (2 * k)) * (fact k) * sqrt (2 / pi)))" by (auto simp: fun_eq_iff power_mult field_simps real_sqrt_power[symmetric] real_sqrt_divide real_sqrt_mult) @@ -1115,7 +1115,7 @@ note normal_moment_odd[OF \_pos, of \ 0] has_bochner_integral_mult_left[of \, OF this] note has_bochner_integral_add[OF this] then show ?thesis - by (simp add: power2_eq_square field_simps) + by (simp add: power2_eq_square field_simps) qed lemma integral_normal_moment_nz_1: @@ -1213,13 +1213,13 @@ assume "distributed M lborel X (\x. ereal (normal_density \ \ x))" then have "distributed M lborel (\x. -\ / \ + (1/\) * X x) (\x. ereal (normal_density (-\ / \ + (1/\)* \) (\1/\\ * \) x))" by(rule normal_density_affine) auto - + then show "distributed M lborel (\x. (X x - \) / \) (\x. ereal (std_normal_density x))" by (simp add: diff_divide_distrib[symmetric] field_simps) next assume *: "distributed M lborel (\x. (X x - \) / \) (\x. ereal (std_normal_density x))" have "distributed M lborel (\x. \ + \ * ((X x - \) / \)) (\x. ereal (normal_density \ \\\ x))" - using normal_density_affine[OF *, of \ \] by simp + using normal_density_affine[OF *, of \ \] by simp then show "distributed M lborel X (\x. ereal (normal_density \ \ x))" by simp qed @@ -1231,8 +1231,8 @@ def \' \ "\\<^sup>2" and \' \ "\\<^sup>2" then have [simp, arith]: "0 < \'" "0 < \'" by simp_all - let ?\ = "sqrt ((\' * \') / (\' + \'))" - have sqrt: "(sqrt (2 * pi * (\' + \')) * sqrt (2 * pi * (\' * \') / (\' + \'))) = + let ?\ = "sqrt ((\' * \') / (\' + \'))" + have sqrt: "(sqrt (2 * pi * (\' + \')) * sqrt (2 * pi * (\' * \') / (\' + \'))) = (sqrt (2 * pi * \') * sqrt (2 * pi * \'))" by (subst power_eq_iff_eq_base[symmetric, where n=2]) (simp_all add: real_sqrt_mult[symmetric] power2_eq_square) @@ -1269,10 +1269,10 @@ have ind[simp]: "indep_var borel (\x. - \ + X x) borel (\x. - \ + Y x)" proof - have "indep_var borel ( (\x. -\ + x) o X) borel ((\x. - \ + x) o Y)" - by (auto intro!: indep_var_compose assms) + by (auto intro!: indep_var_compose assms) then show ?thesis by (simp add: o_def) qed - + have "distributed M lborel (\x. -\ + 1 * X x) (normal_density (- \ + 1 * \) (\1\ * \))" by(rule normal_density_affine[OF normalX pos_var(1), of 1 "-\"]) simp then have 1[simp]: "distributed M lborel (\x. - \ + X x) (normal_density 0 \)" by simp @@ -1280,10 +1280,10 @@ have "distributed M lborel (\x. -\ + 1 * Y x) (normal_density (- \ + 1 * \) (\1\ * \))" by(rule normal_density_affine[OF normalY pos_var(2), of 1 "-\"]) simp then have 2[simp]: "distributed M lborel (\x. - \ + Y x) (normal_density 0 \)" by simp - + have *: "distributed M lborel (\x. (- \ + X x) + (- \ + Y x)) (\x. ereal (normal_density 0 (sqrt (\\<^sup>2 + \\<^sup>2)) x))" using distributed_convolution[OF ind 1 2] conv_normal_density_zero_mean[OF pos_var] by simp - + have "distributed M lborel (\x. \ + \ + 1 * (- \ + X x + (- \ + Y x))) (\x. ereal (normal_density (\ + \ + 1 * 0) (\1\ * sqrt (\\<^sup>2 + \\<^sup>2)) x))" by (rule normal_density_affine[OF *, of 1 "\ + \"]) (auto simp: add_pos_pos) @@ -1298,7 +1298,7 @@ assumes normalY[simp]: "distributed M lborel Y (normal_density \ \)" shows "distributed M lborel (\x. X x - Y x) (normal_density (\ - \) (sqrt (\\<^sup>2 + \\<^sup>2)))" proof - - have "distributed M lborel (\x. 0 + - 1 * Y x) (\x. ereal (normal_density (0 + - 1 * \) (\- 1\ * \) x))" + have "distributed M lborel (\x. 0 + - 1 * Y x) (\x. ereal (normal_density (0 + - 1 * \) (\- 1\ * \) x))" by(rule normal_density_affine, auto) then have [simp]:"distributed M lborel (\x. - Y x) (\x. ereal (normal_density (- \) \ x))" by simp @@ -1315,25 +1315,25 @@ assumes "\i. i \ I \ 0 < \ i" assumes normal: "\i. i \ I \ distributed M lborel (X i) (normal_density (\ i) (\ i))" shows "distributed M lborel (\x. \i\I. X i x) (normal_density (\i\I. \ i) (sqrt (\i\I. (\ i)\<^sup>2)))" - using assms + using assms proof (induct I rule: finite_ne_induct) case (singleton i) then show ?case by (simp add : abs_of_pos) next case (insert i I) - then have 1: "distributed M lborel (\x. (X i x) + (\i\I. X i x)) + then have 1: "distributed M lborel (\x. (X i x) + (\i\I. X i x)) (normal_density (\ i + setsum \ I) (sqrt ((\ i)\<^sup>2 + (sqrt (\i\I. (\ i)\<^sup>2))\<^sup>2)))" - apply (intro sum_indep_normal indep_vars_setsum insert real_sqrt_gt_zero) + apply (intro sum_indep_normal indep_vars_setsum insert real_sqrt_gt_zero) apply (auto intro: indep_vars_subset intro!: setsum_pos) apply fastforce done have 2: "(\x. (X i x)+ (\i\I. X i x)) = (\x. (\j\insert i I. X j x))" "\ i + setsum \ I = setsum \ (insert i I)" using insert by auto - + have 3: "(sqrt ((\ i)\<^sup>2 + (sqrt (\i\I. (\ i)\<^sup>2))\<^sup>2)) = (sqrt (\i\insert i I. (\ i)\<^sup>2))" using insert by (simp add: setsum_nonneg) - - show ?case using 1 2 3 by simp + + show ?case using 1 2 3 by simp qed lemma (in prob_space) standard_normal_distributed_expectation: diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Giry_Monad.thy --- a/src/HOL/Probability/Giry_Monad.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Giry_Monad.thy Tue Nov 10 14:18:41 2015 +0000 @@ -337,7 +337,7 @@ shows "(\M. integral\<^sup>L M f) \ borel_measurable (subprob_algebra N)" proof - note [measurable] = nn_integral_measurable_subprob_algebra - have "?thesis \ (\M. real (\\<^sup>+ x. f x \M) - real (\\<^sup>+ x. - f x \M)) \ borel_measurable (subprob_algebra N)" + have "?thesis \ (\M. real_of_ereal (\\<^sup>+ x. f x \M) - real_of_ereal (\\<^sup>+ x. - f x \M)) \ borel_measurable (subprob_algebra N)" proof(rule measurable_cong) fix M assume "M \ space (subprob_algebra N)" @@ -352,7 +352,7 @@ finally have "(\\<^sup>+ x. ereal \f x\ \M) \ \" by(auto simp add: max_def) then have "integrable M f" using f_measurable by(auto intro: integrableI_bounded) - thus "(\ x. f x \M) = real (\\<^sup>+ x. f x \M) - real (\\<^sup>+ x. - f x \M)" + thus "(\ x. f x \M) = real_of_ereal (\\<^sup>+ x. f x \M) - real_of_ereal (\\<^sup>+ x. - f x \M)" by(simp add: real_lebesgue_integral_def) qed also have "\" by measurable @@ -933,9 +933,9 @@ using f_measurable by(auto intro!: bounded1 dest: f_bounded) have "AE M' in M. (\\<^sup>+ x. f x \M') \ \" using M_bounded by(rule AE_mp[OF _ AE_I2])(auto dest: f_pos1) - hence [simp]: "(\\<^sup>+ M'. ereal (real (\\<^sup>+ x. f x \M')) \M) = (\\<^sup>+ M'. \\<^sup>+ x. f x \M' \M)" + hence [simp]: "(\\<^sup>+ M'. ereal (real_of_ereal (\\<^sup>+ x. f x \M')) \M) = (\\<^sup>+ M'. \\<^sup>+ x. f x \M' \M)" by(rule nn_integral_cong_AE[OF AE_mp])(simp add: ereal_real nn_integral_nonneg) - from f_pos have [simp]: "integrable M (\M'. real (\\<^sup>+ x. f x \M'))" + from f_pos have [simp]: "integrable M (\M'. real_of_ereal (\\<^sup>+ x. f x \M'))" by(simp add: int_f real_integrable_def nn_integral_nonneg real_of_ereal[symmetric] nn_integral_0_iff_AE[THEN iffD2] del: real_of_ereal) have f_neg1: @@ -944,9 +944,9 @@ using f_measurable by(auto intro!: bounded1 dest: f_bounded) have "AE M' in M. (\\<^sup>+ x. - f x \M') \ \" using M_bounded by(rule AE_mp[OF _ AE_I2])(auto dest: f_neg1) - hence [simp]: "(\\<^sup>+ M'. ereal (real (\\<^sup>+ x. - f x \M')) \M) = (\\<^sup>+ M'. \\<^sup>+ x. - f x \M' \M)" + hence [simp]: "(\\<^sup>+ M'. ereal (real_of_ereal (\\<^sup>+ x. - f x \M')) \M) = (\\<^sup>+ M'. \\<^sup>+ x. - f x \M' \M)" by(rule nn_integral_cong_AE[OF AE_mp])(simp add: ereal_real nn_integral_nonneg) - from f_neg have [simp]: "integrable M (\M'. real (\\<^sup>+ x. - f x \M'))" + from f_neg have [simp]: "integrable M (\M'. real_of_ereal (\\<^sup>+ x. - f x \M'))" by(simp add: int_mf real_integrable_def nn_integral_nonneg real_of_ereal[symmetric] nn_integral_0_iff_AE[THEN iffD2] del: real_of_ereal) from \?integrable\ @@ -958,15 +958,15 @@ ((\\<^sup>+ M'. \\<^sup>+ x. f x \M' \M) - (\\<^sup>+ M'. - \\<^sup>+ x. f x \M' \M)) - ((\\<^sup>+ M'. \\<^sup>+ x. - f x \M' \M) - (\\<^sup>+ M'. - \\<^sup>+ x. - f x \M' \M))" by(subst (7 11) nn_integral_0_iff_AE[THEN iffD2])(simp_all add: nn_integral_nonneg) - also have "(\\<^sup>+ M'. \\<^sup>+ x. f x \M' \M) - (\\<^sup>+ M'. - \\<^sup>+ x. f x \M' \M) = \ M'. real (\\<^sup>+ x. f x \M') \M" + also have "(\\<^sup>+ M'. \\<^sup>+ x. f x \M' \M) - (\\<^sup>+ M'. - \\<^sup>+ x. f x \M' \M) = \ M'. real_of_ereal (\\<^sup>+ x. f x \M') \M" using f_pos by(simp add: real_lebesgue_integral_def)(simp add: ereal_minus(1)[symmetric] ereal_real int_f nn_integral_nonneg nn_integral_0_iff_AE[THEN iffD2] real_of_ereal_pos zero_ereal_def[symmetric]) - also have "(\\<^sup>+ M'. \\<^sup>+ x. - f x \M' \M) - (\\<^sup>+ M'. - \\<^sup>+ x. - f x \M' \M) = \ M'. real (\\<^sup>+ x. - f x \M') \M" + also have "(\\<^sup>+ M'. \\<^sup>+ x. - f x \M' \M) - (\\<^sup>+ M'. - \\<^sup>+ x. - f x \M' \M) = \ M'. real_of_ereal (\\<^sup>+ x. - f x \M') \M" using f_neg by(simp add: real_lebesgue_integral_def)(simp add: ereal_minus(1)[symmetric] ereal_real int_mf nn_integral_nonneg nn_integral_0_iff_AE[THEN iffD2] real_of_ereal_pos zero_ereal_def[symmetric]) also note ereal_minus(1) - also have "(\ M'. real (\\<^sup>+ x. f x \M') \M) - (\ M'. real (\\<^sup>+ x. - f x \M') \M) = - \M'. real (\\<^sup>+ x. f x \M') - real (\\<^sup>+ x. - f x \M') \M" + also have "(\ M'. real_of_ereal (\\<^sup>+ x. f x \M') \M) - (\ M'. real_of_ereal (\\<^sup>+ x. - f x \M') \M) = + \M'. real_of_ereal (\\<^sup>+ x. f x \M') - real_of_ereal (\\<^sup>+ x. - f x \M') \M" by simp also have "\ = \M'. \ x. f x \M' \M" using _ _ M_bounded proof(rule integral_cong_AE[OF _ _ AE_mp[OF _ AE_I2], rule_format]) @@ -982,7 +982,7 @@ hence [simp]: "space M' = space N" by(rule sets_eq_imp_space_eq) have "integrable M' f" by(rule integrable_const_bound[where B=B])(auto simp add: f_bounded) - then show "real (\\<^sup>+ x. f x \M') - real (\\<^sup>+ x. - f x \M') = \ x. f x \M'" + then show "real_of_ereal (\\<^sup>+ x. f x \M') - real_of_ereal (\\<^sup>+ x. - f x \M') = \ x. f x \M'" by(simp add: real_lebesgue_integral_def) qed simp_all finally show ?integral by simp diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Information.thy Tue Nov 10 14:18:41 2015 +0000 @@ -73,7 +73,7 @@ Kullback$-$Leibler distance. *} definition - "entropy_density b M N = log b \ real \ RN_deriv M N" + "entropy_density b M N = log b \ real_of_ereal \ RN_deriv M N" definition "KL_divergence b M N = integral\<^sup>L N (entropy_density b M N)" @@ -247,7 +247,7 @@ apply auto done qed - then have "AE x in M. log b (real (RN_deriv M M x)) = 0" + then have "AE x in M. log b (real_of_ereal (RN_deriv M M x)) = 0" by (elim AE_mp) simp from integral_cong_AE[OF _ _ this] have "integral\<^sup>L M (entropy_density b M M) = 0" @@ -787,7 +787,7 @@ note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X] note ae = distributed_RN_deriv[OF X] - have ae_eq: "AE x in distr M MX X. log b (real (RN_deriv MX (distr M MX X) x)) = + have ae_eq: "AE x in distr M MX X. log b (real_of_ereal (RN_deriv MX (distr M MX X) x)) = log b (f x)" unfolding distributed_distr_eq_density[OF X] apply (subst AE_density) @@ -1517,8 +1517,8 @@ subsection {* Conditional Entropy *} definition (in prob_space) - "conditional_entropy b S T X Y = - (\(x, y). log b (real (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) (x, y)) / - real (RN_deriv T (distr M T Y) y)) \distr M (S \\<^sub>M T) (\x. (X x, Y x)))" + "conditional_entropy b S T X Y = - (\(x, y). log b (real_of_ereal (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) (x, y)) / + real_of_ereal (RN_deriv T (distr M T Y) y)) \distr M (S \\<^sub>M T) (\x. (X x, Y x)))" abbreviation (in information_space) conditional_entropy_Pow ("\'(_ | _')") where @@ -1535,13 +1535,13 @@ interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. - have "AE x in density (S \\<^sub>M T) (\x. ereal (Pxy x)). Pxy x = real (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) x)" + have "AE x in density (S \\<^sub>M T) (\x. ereal (Pxy x)). Pxy x = real_of_ereal (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) x)" unfolding AE_density[OF distributed_borel_measurable, OF Pxy] unfolding distributed_distr_eq_density[OF Pxy] using distributed_RN_deriv[OF Pxy] by auto moreover - have "AE x in density (S \\<^sub>M T) (\x. ereal (Pxy x)). Py (snd x) = real (RN_deriv T (distr M T Y) (snd x))" + have "AE x in density (S \\<^sub>M T) (\x. ereal (Pxy x)). Py (snd x) = real_of_ereal (RN_deriv T (distr M T Y) (snd x))" unfolding AE_density[OF distributed_borel_measurable, OF Pxy] unfolding distributed_distr_eq_density[OF Py] apply (rule ST.AE_pair_measure) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Interval_Integral.thy --- a/src/HOL/Probability/Interval_Integral.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Interval_Integral.thy Tue Nov 10 14:18:41 2015 +0000 @@ -67,19 +67,21 @@ case PInf with `a < b` have "a = -\ \ (\r. a = ereal r)" by (cases a) auto - moreover have " (\x. ereal (real (Suc x))) ----> \" - using nat_ceiling_le_eq by (subst LIMSEQ_Suc_iff) (auto simp: Lim_PInfty) + moreover have "(\x. ereal (real (Suc x))) ----> \" + apply (subst LIMSEQ_Suc_iff) + apply (simp add: Lim_PInfty) + using nat_ceiling_le_eq by blast moreover have "\r. (\x. ereal (r + real (Suc x))) ----> \" apply (subst LIMSEQ_Suc_iff) apply (subst Lim_PInfty) apply (metis add.commute diff_le_eq nat_ceiling_le_eq ereal_less_eq(3)) done ultimately show thesis - by (intro that[of "\i. real a + Suc i"]) + by (intro that[of "\i. real_of_ereal a + Suc i"]) (auto simp: incseq_def PInf) next case (real b') - def d \ "b' - (if a = -\ then b' - 1 else real a)" + def d \ "b' - (if a = -\ then b' - 1 else real_of_ereal a)" with `a < b` have a': "0 < d" by (cases a) (auto simp: real) moreover @@ -367,7 +369,7 @@ lemma interval_integral_Icc': "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {x. a \ ereal x \ ereal x \ b}. f x)" - by (auto intro!: set_integral_discrete_difference[where X="{real a, real b}"] + by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"] simp add: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ioc: @@ -378,7 +380,7 @@ (* TODO: other versions as well? *) (* Yes: I need the Icc' version. *) lemma interval_integral_Ioc': "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {x. a < ereal x \ ereal x \ b}. f x)" - by (auto intro!: set_integral_discrete_difference[where X="{real a, real b}"] + by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"] simp add: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ico: @@ -420,7 +422,7 @@ have "(LBINT x:einterval a c. f x) = (LBINT x:einterval a b \ einterval b c. f x)" proof (rule set_integral_cong_set) show "AE x in lborel. (x \ einterval a b \ einterval b c) = (x \ einterval a c)" - using AE_lborel_singleton[of "real b"] ord + using AE_lborel_singleton[of "real_of_ereal b"] ord by (cases a b c rule: ereal3_cases) (auto simp: einterval_iff) qed (insert ord, auto intro!: set_borel_measurable_subset[OF f] simp: einterval_iff) also have "\ = (LBINT x:einterval a b. f x) + (LBINT x:einterval b c. f x)" @@ -607,8 +609,8 @@ assumes F: "\x. a < ereal x \ ereal x < b \ DERIV F x :> f x" assumes f: "\x. a < ereal x \ ereal x < b \ isCont f x" assumes f_nonneg: "AE x in lborel. a < ereal x \ ereal x < b \ 0 \ f x" - assumes A: "((F \ real) ---> A) (at_right a)" - assumes B: "((F \ real) ---> B) (at_left b)" + assumes A: "((F \ real_of_ereal) ---> A) (at_right a)" + assumes B: "((F \ real_of_ereal) ---> B) (at_left b)" shows "set_integrable lborel (einterval a b) f" "(LBINT x=a..b. f x) = B - A" @@ -654,8 +656,8 @@ assumes F: "\x. a < ereal x \ ereal x < b \ (F has_vector_derivative f x) (at x)" assumes f: "\x. a < ereal x \ ereal x < b \ isCont f x" assumes f_integrable: "set_integrable lborel (einterval a b) f" - assumes A: "((F \ real) ---> A) (at_right a)" - assumes B: "((F \ real) ---> B) (at_left b)" + assumes A: "((F \ real_of_ereal) ---> A) (at_right a)" + assumes B: "((F \ real_of_ereal) ---> B) (at_left b)" shows "(LBINT x=a..b. f x) = B - A" proof - from einterval_Icc_approximation[OF `a < b`] guess u l . note approx = this @@ -819,8 +821,8 @@ and contf: "\x. a < ereal x \ ereal x < b \ isCont f (g x)" and contg': "\x. a < ereal x \ ereal x < b \ isCont g' x" and g'_nonneg: "\x. a \ ereal x \ ereal x \ b \ 0 \ g' x" - and A: "((ereal \ g \ real) ---> A) (at_right a)" - and B: "((ereal \ g \ real) ---> B) (at_left b)" + and A: "((ereal \ g \ real_of_ereal) ---> A) (at_right a)" + and B: "((ereal \ g \ real_of_ereal) ---> B) (at_left b)" and integrable: "set_integrable lborel (einterval a b) (\x. g' x *\<^sub>R f (g x))" and integrable2: "set_integrable lborel (einterval A B) (\x. f x)" shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))" @@ -920,8 +922,8 @@ and contg': "\x. a < ereal x \ ereal x < b \ isCont g' x" and f_nonneg: "\x. a < ereal x \ ereal x < b \ 0 \ f (g x)" (* TODO: make this AE? *) and g'_nonneg: "\x. a \ ereal x \ ereal x \ b \ 0 \ g' x" - and A: "((ereal \ g \ real) ---> A) (at_right a)" - and B: "((ereal \ g \ real) ---> B) (at_left b)" + and A: "((ereal \ g \ real_of_ereal) ---> A) (at_right a)" + and B: "((ereal \ g \ real_of_ereal) ---> B) (at_left b)" and integrable_fg: "set_integrable lborel (einterval a b) (\x. f (g x) * g' x)" shows "set_integrable lborel (einterval A B) f" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Lebesgue_Integral_Substitution.thy --- a/src/HOL/Probability/Lebesgue_Integral_Substitution.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Lebesgue_Integral_Substitution.thy Tue Nov 10 14:18:41 2015 +0000 @@ -327,7 +327,7 @@ from derivg have derivg': "\x. x \ {a..b} \ (g has_vector_derivative g' x) (at x)" by (simp only: has_field_derivative_iff_has_vector_derivative) - have real_ind[simp]: "\A x. real (indicator A x :: ereal) = indicator A x" + have real_ind[simp]: "\A x. real_of_ereal (indicator A x :: ereal) = indicator A x" by (auto split: split_indicator) have ereal_ind[simp]: "\A x. ereal (indicator A x) = indicator A x" by (auto split: split_indicator) @@ -645,8 +645,8 @@ unfolding real_integrable_def by (force simp: mult.commute) have "LBINT x. (f x :: real) * indicator {g a..g b} x = - real (\\<^sup>+ x. ereal (f x) * indicator {g a..g b} x \lborel) - - real (\\<^sup>+ x. ereal (- (f x)) * indicator {g a..g b} x \lborel)" using integrable + real_of_ereal (\\<^sup>+ x. ereal (f x) * indicator {g a..g b} x \lborel) - + real_of_ereal (\\<^sup>+ x. ereal (- (f x)) * indicator {g a..g b} x \lborel)" using integrable by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ereal mult.commute) also have "(\\<^sup>+x. ereal (f x) * indicator {g a..g b} x \lborel) = (\\<^sup>+x. ereal (f x * indicator {g a..g b} x) \lborel)" @@ -677,8 +677,8 @@ by (simp_all add: nn_integral_set_ereal mult.commute) } note integrable' = this - have "real (\\<^sup>+ x. ereal (f (g x) * g' x * indicator {a..b} x) \lborel) - - real (\\<^sup>+ x. ereal (-f (g x) * g' x * indicator {a..b} x) \lborel) = + have "real_of_ereal (\\<^sup>+ x. ereal (f (g x) * g' x * indicator {a..b} x) \lborel) - + real_of_ereal (\\<^sup>+ x. ereal (-f (g x) * g' x * indicator {a..b} x) \lborel) = (LBINT x. f (g x) * g' x * indicator {a..b} x)" using integrable' by (subst real_lebesgue_integral_def) (simp_all add: field_simps) finally show "(LBINT x. f x * indicator {g a..g b} x) = diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Nov 10 14:18:41 2015 +0000 @@ -510,7 +510,7 @@ also have "... \ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" apply (rule mult_left_mono) - apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le real_of_nat_le_iff real_of_nat_power self_le_power zero_less_Suc) + apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc) apply (simp add: DIM_positive) done finally have "real n \ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" . @@ -963,7 +963,7 @@ note F(1)[THEN borel_measurable_simple_function, measurable] - { fix i x have "real (F i x) \ f x" + { fix i x have "real_of_ereal (F i x) \ f x" using F(3,5) F(4)[of x, symmetric] nonneg unfolding real_le_ereal_iff by (auto simp: image_iff eq_commute[of \] max_def intro: SUP_upper split: split_if_asm) } @@ -991,25 +991,25 @@ also have "\ \ ereal I" proof (rule SUP_least) fix i :: nat - have finite_F: "(\\<^sup>+ x. ereal (real (F i x) * indicator (?B i) x) \lborel) < \" + have finite_F: "(\\<^sup>+ x. ereal (real_of_ereal (F i x) * indicator (?B i) x) \lborel) < \" proof (rule nn_integral_bound_simple_function) - have "emeasure lborel {x \ space lborel. ereal (real (F i x) * indicator (?B i) x) \ 0} \ + have "emeasure lborel {x \ space lborel. ereal (real_of_ereal (F i x) * indicator (?B i) x) \ 0} \ emeasure lborel (?B i)" by (intro emeasure_mono) (auto split: split_indicator) - then show "emeasure lborel {x \ space lborel. ereal (real (F i x) * indicator (?B i) x) \ 0} < \" + then show "emeasure lborel {x \ space lborel. ereal (real_of_ereal (F i x) * indicator (?B i) x) \ 0} < \" by auto qed (auto split: split_indicator - intro!: real_of_ereal_pos F simple_function_compose1[where g="real"] simple_function_ereal) + intro!: real_of_ereal_pos F simple_function_compose1[where g="real_of_ereal"] simple_function_ereal) - have int_F: "(\x. real (F i x) * indicator (?B i) x) integrable_on UNIV" + have int_F: "(\x. real_of_ereal (F i x) * indicator (?B i) x) integrable_on UNIV" using F(5) finite_F by (intro nn_integral_integrable_on) (auto split: split_indicator intro: real_of_ereal_pos) have "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) = - (\\<^sup>+ x. ereal (real (F i x) * indicator (?B i) x) \lborel)" + (\\<^sup>+ x. ereal (real_of_ereal (F i x) * indicator (?B i) x) \lborel)" using F(3,5) by (intro nn_integral_cong) (auto simp: image_iff ereal_real eq_commute split: split_indicator) - also have "\ = ereal (integral UNIV (\x. real (F i x) * indicator (?B i) x))" + also have "\ = ereal (integral UNIV (\x. real_of_ereal (F i x) * indicator (?B i) x))" using F by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F]) (auto split: split_indicator intro: real_of_ereal_pos) @@ -1279,7 +1279,7 @@ fix m n :: nat assume "m \ n" with f nonneg show "F (a + real m) \ F (a + real n)" by (intro DERIV_nonneg_imp_nondecreasing[where f=F]) - (simp, metis add_increasing2 order_refl order_trans real_of_nat_ge_zero) + (simp, metis add_increasing2 order_refl order_trans of_nat_0_le_iff) qed have "(\x. F (a + real x)) ----> T" apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top]) @@ -1298,7 +1298,7 @@ proof (subst integral_FTC_Icc_real) fix x show "DERIV (\x. x^Suc k / Suc k) x :> x^k" by (intro derivative_eq_intros) auto -qed (auto simp: field_simps simp del: real_of_nat_Suc) +qed (auto simp: field_simps simp del: of_nat_Suc) subsection {* Integration by parts *} diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Measure_Space.thy --- a/src/HOL/Probability/Measure_Space.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Measure_Space.thy Tue Nov 10 14:18:41 2015 +0000 @@ -700,7 +700,7 @@ lemma emeasure_insert_ne: "A \ {} \ {x} \ sets M \ A \ sets M \ x \ A \ emeasure M (insert x A) = emeasure M {x} + emeasure M A" - by (rule emeasure_insert) + by (rule emeasure_insert) lemma emeasure_eq_setsum_singleton: assumes "finite S" "\x. x \ S \ {x} \ sets M" @@ -877,7 +877,7 @@ by (auto intro!: antisym) qed -lemma UN_from_nat_into: +lemma UN_from_nat_into: assumes I: "countable I" "I \ {}" shows "(\i\I. N i) = (\i. N (from_nat_into I i))" proof - @@ -1100,7 +1100,7 @@ unfolding eventually_ae_filter by auto qed auto -lemma AE_ball_countable: +lemma AE_ball_countable: assumes [intro]: "countable X" shows "(AE x in M. \y\X. P x y) \ (\y\X. AE x in M. P x y)" proof @@ -1121,7 +1121,7 @@ lemma AE_discrete_difference: assumes X: "countable X" - assumes null: "\x. x \ X \ emeasure M {x} = 0" + assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" shows "AE x in M. x \ X" proof - @@ -1207,7 +1207,7 @@ assume "A = {}" with `(\A) = space M` show thesis by (intro that[of "\_. {}"]) auto next - assume "A \ {}" + assume "A \ {}" show thesis proof show "range (from_nat_into A) \ sets M" @@ -1364,7 +1364,7 @@ by (simp add: emeasure_distr measure_def) lemma distr_cong_AE: - assumes 1: "M = K" "sets N = sets L" and + assumes 1: "M = K" "sets N = sets L" and 2: "(AE x in M. f x = g x)" and "f \ measurable M N" and "g \ measurable K L" shows "distr M N f = distr K L g" proof (rule measure_eqI) @@ -1713,7 +1713,7 @@ using finite_measure_eq_setsum_singleton[OF s] by simp also have "\ = (\ x \ s. measure M {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * measure M {(SOME x. x \ s)}" - using setsum_constant assms by (simp add: real_eq_of_nat) + using setsum_constant assms by simp finally show ?thesis by simp qed simp @@ -1744,7 +1744,7 @@ shows "measure M B = 0" using measure_space_inter[of B A] assms by (auto simp: ac_simps) lemma (in finite_measure) finite_measure_distr: - assumes f: "f \ measurable M M'" + assumes f: "f \ measurable M M'" shows "finite_measure (distr M M' f)" proof (rule finite_measureI) have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) @@ -1795,7 +1795,7 @@ interpret ring_of_sets A "Pow A" by (rule ring_of_setsI) auto - show "countably_additive (Pow A) ?M" + show "countably_additive (Pow A) ?M" unfolding countably_additive_iff_continuous_from_below[OF positive additive] proof safe fix F :: "nat \ 'a set" assume "incseq F" @@ -1844,7 +1844,7 @@ by (rule infinite_super[OF _ 1]) auto then have "infinite (\i. F i)" by auto - + ultimately show ?thesis by auto qed qed @@ -1975,7 +1975,7 @@ by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff emeasure_restrict_space cong: conj_cong intro!: ex_cong[where f="\X. (\ \ space M) \ X"]) -qed +qed lemma restrict_restrict_space: assumes "A \ space M \ sets M" "B \ space M \ sets M" @@ -2036,7 +2036,7 @@ by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space) qed -lemma restrict_distr: +lemma restrict_distr: assumes [measurable]: "f \ measurable M N" assumes [simp]: "\ \ space N \ sets N" and restrict: "f \ space M \ \" shows "restrict_space (distr M N f) \ = distr M (restrict_space N \) f" @@ -2053,7 +2053,7 @@ assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" assumes sets_eq: "sets M = sets N" and \: "\ \ sets M" assumes "sets (restrict_space M \) = sigma_sets \ E" - assumes "sets (restrict_space N \) = sigma_sets \ E" + assumes "sets (restrict_space N \) = sigma_sets \ E" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" assumes A: "countable A" "A \ {}" "A \ E" "\A = \" "\a. a \ A \ emeasure M a \ \" shows "M = N" @@ -2089,7 +2089,7 @@ lemma space_null_measure[simp]: "space (null_measure M) = space M" by (simp add: null_measure_def) -lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" +lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" by (simp add: null_measure_def) lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0" @@ -2176,12 +2176,12 @@ proof - have sets: "(SUP a:A. sets a) = sets a" proof (intro antisym SUP_least) - fix a' show "a' \ A \ sets a' \ sets a" + fix a' show "a' \ A \ sets a' \ sets a" using a chainD[OF A, of a a'] by (auto elim!: less_eq_measure.cases) qed (insert \a\A\, auto) have space: "(SUP a:A. space a) = space a" proof (intro antisym SUP_least) - fix a' show "a' \ A \ space a' \ space a" + fix a' show "a' \ A \ space a' \ space a" using a chainD[OF A, of a a'] by (intro sets_le_imp_space_le) (auto elim!: less_eq_measure.cases) qed (insert \a\A\, auto) show "space (Sup A) = space a" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Tue Nov 10 14:18:41 2015 +0000 @@ -204,12 +204,12 @@ assume "finite P" from this assms show ?thesis by induct auto qed auto -lemma simple_function_ereal[intro, simp]: +lemma simple_function_ereal[intro, simp]: fixes f g :: "'a \ real" assumes sf: "simple_function M f" shows "simple_function M (\x. ereal (f x))" by (rule simple_function_compose1[OF sf]) -lemma simple_function_real_of_nat[intro, simp]: +lemma simple_function_real_of_nat[intro, simp]: fixes f g :: "'a \ nat" assumes sf: "simple_function M f" shows "simple_function M (\x. real (f x))" by (rule simple_function_compose1[OF sf]) @@ -220,21 +220,21 @@ shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ (\x. (SUP i. f i x) = max 0 (u x)) \ (\i x. 0 \ f i x)" proof - - def f \ "\x i. if real i \ u x then i * 2 ^ i else nat(floor (real (u x) * 2 ^ i))" + def f \ "\x i. if real i \ u x then i * 2 ^ i else nat(floor (real_of_ereal (u x) * 2 ^ i))" { fix x j have "f x j \ j * 2 ^ j" unfolding f_def proof (split split_if, intro conjI impI) assume "\ real j \ u x" - then have "nat(floor (real (u x) * 2 ^ j)) \ nat(floor (j * 2 ^ j))" + then have "nat(floor (real_of_ereal (u x) * 2 ^ j)) \ nat(floor (j * 2 ^ j))" by (cases "u x") (auto intro!: nat_mono floor_mono) moreover have "real (nat(floor (j * 2 ^ j))) \ j * 2^j" by linarith - ultimately show "nat(floor (real (u x) * 2 ^ j)) \ j * 2 ^ j" - unfolding real_of_nat_le_iff by auto + ultimately show "nat(floor (real_of_ereal (u x) * 2 ^ j)) \ j * 2 ^ j" + unfolding of_nat_le_iff by auto qed auto } note f_upper = this have real_f: - "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (nat(floor (real (u x) * 2 ^ i))))" + "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (nat(floor (real_of_ereal (u x) * 2 ^ i))))" unfolding f_def by auto let ?g = "\j x. real (f x j) / 2^j :: ereal" @@ -259,27 +259,27 @@ have "f x i * 2 \ f x (Suc i)" unfolding f_def proof ((split split_if)+, intro conjI impI) assume "ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" - then show "i * 2 ^ i * 2 \ nat(floor (real (u x) * 2 ^ Suc i))" + then show "i * 2 ^ i * 2 \ nat(floor (real_of_ereal (u x) * 2 ^ Suc i))" by (cases "u x") (auto intro!: le_nat_floor) next assume "\ ereal (real i) \ u x" "ereal (real (Suc i)) \ u x" - then show "nat(floor (real (u x) * 2 ^ i)) * 2 \ Suc i * 2 ^ Suc i" + then show "nat(floor (real_of_ereal (u x) * 2 ^ i)) * 2 \ Suc i * 2 ^ Suc i" by (cases "u x") auto next assume "\ ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" - have "nat(floor (real (u x) * 2 ^ i)) * 2 = nat(floor (real (u x) * 2 ^ i)) * nat(floor(2::real))" + have "nat(floor (real_of_ereal (u x) * 2 ^ i)) * 2 = nat(floor (real_of_ereal (u x) * 2 ^ i)) * nat(floor(2::real))" by simp - also have "\ \ nat(floor (real (u x) * 2 ^ i * 2))" + also have "\ \ nat(floor (real_of_ereal (u x) * 2 ^ i * 2))" proof cases assume "0 \ u x" then show ?thesis - by (intro le_mult_nat_floor) + by (intro le_mult_nat_floor) next assume "\ 0 \ u x" then show ?thesis by (cases "u x") (auto simp: nat_floor_neg mult_nonpos_nonneg) qed - also have "\ = nat(floor (real (u x) * 2 ^ Suc i))" + also have "\ = nat(floor (real_of_ereal (u x) * 2 ^ Suc i))" by (simp add: ac_simps) - finally show "nat(floor (real (u x) * 2 ^ i)) * 2 \ nat(floor (real (u x) * 2 ^ Suc i))" . + finally show "nat(floor (real_of_ereal (u x) * 2 ^ i)) * 2 \ nat(floor (real_of_ereal (u x) * 2 ^ Suc i))" . qed simp then show "?g i x \ ?g (Suc i) x" by (auto simp: field_simps) @@ -380,7 +380,7 @@ apply (auto intro: u) done next - + from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" unfolding simple_function_def by auto then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" @@ -417,7 +417,7 @@ have not_inf: "\x i. x \ space M \ U i x < \" using U by (auto simp: image_iff eq_commute) - + from U have "\i. U i \ borel_measurable M" by (simp add: borel_measurable_simple_function) @@ -556,12 +556,12 @@ note eq = this have "integral\<^sup>S M f = - (\y\f`space M. y * (\z\g`space M. + (\y\f`space M. y * (\z\g`space M. if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" unfolding simple_integral_def proof (safe intro!: setsum.cong ereal_right_mult_cong) fix y assume y: "y \ space M" "f y \ 0" - have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = + have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = {z. \x\space M. f y = f x \ z = g x}" by auto have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = @@ -577,7 +577,7 @@ apply (auto simp: disjoint_family_on_def eq) done qed - also have "\ = (\y\f`space M. (\z\g`space M. + also have "\ = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" by (auto intro!: setsum.cong simp: setsum_ereal_right_distrib emeasure_nonneg) also have "\ = ?r" @@ -1239,7 +1239,7 @@ assume "space M \ {}" with simple_functionD(1)[OF f] bnd have bnd: "0 \ Max (f`space M) \ Max (f`space M) < \" by (subst Max_less_iff) (auto simp: Max_ge_iff) - + have "nn_integral M f \ (\\<^sup>+x. Max (f`space M) * indicator {x\space M. f x \ 0} x \M)" proof (rule nn_integral_mono) fix x assume "x \ space M" @@ -1532,7 +1532,7 @@ by (intro nn_integral_cong) (simp add: sup_ereal_def[symmetric] sup_INF del: sup_ereal_def) also have "\ = (\\<^sup>+ x. (INF j. max 0 (restrict (f (j + i)) (space M) x)) \M)" - using f by (intro nn_integral_cong INF_shift antimonoI le_funI max.mono) + using f by (intro nn_integral_cong INF_shift antimonoI le_funI max.mono) (auto simp: decseq_def le_fun_def) also have "\ = (INF j. (\\<^sup>+ x. max 0 (restrict (f (j + i)) (space M) x) \M))" proof (rule nn_integral_monotone_convergence_INF') @@ -1542,7 +1542,7 @@ using fin by (simp add: nn_integral_max_0 cong: nn_integral_cong) qed (intro max.cobounded1 dec f)+ also have "\ = (INF j. (\\<^sup>+ x. max 0 (restrict (f j) (space M) x) \M))" - using f by (intro INF_shift[symmetric] nn_integral_mono antimonoI le_funI max.mono) + using f by (intro INF_shift[symmetric] nn_integral_mono antimonoI le_funI max.mono) (auto simp: decseq_def le_fun_def) finally show ?thesis unfolding nn_integral_max_0 by (simp cong: nn_integral_cong) qed @@ -1670,7 +1670,7 @@ finally show ?thesis . qed -lemma AE_iff_nn_integral: +lemma AE_iff_nn_integral: "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>N M (indicator {x. \ P x}) = 0" by (subst nn_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def sets.sets_Collect_neg indicator_def[abs_def] measurable_If) @@ -1800,7 +1800,7 @@ subsection {* Integral under concrete measures *} -lemma nn_integral_empty: +lemma nn_integral_empty: assumes "space M = {}" shows "nn_integral M f = 0" proof - @@ -1815,7 +1815,7 @@ assumes T: "T \ measurable M M'" and f: "f \ borel_measurable (distr M M' T)" "\x. 0 \ f x" shows "integral\<^sup>N (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" - using f + using f proof induct case (cong f g) with T show ?case @@ -1943,12 +1943,12 @@ qed lemma emeasure_UN_countable: - assumes sets[measurable]: "\i. i \ I \ X i \ sets M" and I[simp]: "countable I" + assumes sets[measurable]: "\i. i \ I \ X i \ sets M" and I[simp]: "countable I" assumes disj: "disjoint_family_on X I" shows "emeasure M (UNION I X) = (\\<^sup>+i. emeasure M (X i) \count_space I)" proof - have eq: "\x. indicator (UNION I X) x = \\<^sup>+ i. indicator (X i) x \count_space I" - proof cases + proof cases fix x assume x: "x \ UNION I X" then obtain j where j: "x \ X j" "j \ I" by auto @@ -2210,7 +2210,7 @@ intro!: setsum.mono_neutral_cong_left ereal_right_mult_cong[OF refl] arg_cong2[where f=emeasure]) lemma one_not_less_zero[simp]: "\ 1 < (0::ereal)" - by (simp add: zero_ereal_def one_ereal_def) + by (simp add: zero_ereal_def one_ereal_def) lemma nn_integral_restrict_space: assumes \[simp]: "\ \ space M \ sets M" @@ -2245,7 +2245,7 @@ by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \ sf]) show "\x. s x \ {0..<\}" using s by (auto simp: image_subset_iff) - from s show "s \ max 0 \ f" + from s show "s \ max 0 \ f" by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) qed then show ?thesis @@ -2281,7 +2281,7 @@ definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" -lemma +lemma shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M" and space_density[simp]: "space (density M f) = space M" by (auto simp: density_def) @@ -2290,7 +2290,7 @@ lemma space_density_imp[measurable_dest]: "\x M f. x \ space (density M f) \ x \ space M" by auto -lemma +lemma shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" @@ -2353,7 +2353,7 @@ apply (intro nn_integral_cong) apply (auto simp: indicator_def) done - have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ + have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" unfolding eq using f `A \ sets M` @@ -2429,7 +2429,7 @@ qed lemma nn_integral_density: - "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ + "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ integral\<^sup>N (density M f) g' = (\\<^sup>+ x. f x * g' x \M)" by (subst (1 2) nn_integral_max_0[symmetric]) (auto intro!: nn_integral_cong_AE @@ -2468,7 +2468,7 @@ by (auto simp: nn_integral_cmult_indicator emeasure_density) lemma measure_density_const: - "A \ sets M \ 0 \ c \ c \ \ \ measure (density M (\_. c)) A = real c * measure M A" + "A \ sets M \ 0 \ c \ c \ \ \ measure (density M (\_. c)) A = real_of_ereal c * measure M A" by (auto simp: emeasure_density_const measure_def) lemma density_density_eq: @@ -2509,12 +2509,12 @@ by (intro measure_eqI) (auto simp: emeasure_density) lemma emeasure_density_add: - assumes X: "X \ sets M" + assumes X: "X \ sets M" assumes Mf[measurable]: "f \ borel_measurable M" assumes Mg[measurable]: "g \ borel_measurable M" assumes nonnegf: "\x. x \ space M \ f x \ 0" assumes nonnegg: "\x. x \ space M \ g x \ 0" - shows "emeasure (density M f) X + emeasure (density M g) X = + shows "emeasure (density M f) X + emeasure (density M g) X = emeasure (density M (\x. f x + g x)) X" using assms apply (subst (1 2 3) emeasure_density, simp_all) [] @@ -2697,8 +2697,8 @@ subsubsection {* Uniform count measure *} definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" - -lemma + +lemma shows space_uniform_count_measure: "space (uniform_count_measure A) = A" and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) @@ -2706,13 +2706,13 @@ lemma sets_uniform_count_measure_count_space[measurable_cong]: "sets (uniform_count_measure A) = sets (count_space A)" by (simp add: sets_uniform_count_measure) - + lemma emeasure_uniform_count_measure: "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" - by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) - + by (simp add: emeasure_point_measure_finite uniform_count_measure_def) + lemma measure_uniform_count_measure: "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" - by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) + by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def) end diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Tue Nov 10 14:18:41 2015 +0000 @@ -59,7 +59,7 @@ note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" using `?M \ 0` - by (simp add: `card X = Suc (Suc n)` real_eq_of_nat[symmetric] real_of_nat_Suc field_simps less_le measure_nonneg) + by (simp add: `card X = Suc (Suc n)` of_nat_Suc field_simps less_le measure_nonneg) also have "\ \ (\x\X. ?m x)" by (rule setsum_mono) fact also have "\ = measure M (\x\X. {x})" @@ -956,7 +956,7 @@ lemma quotient_rel_set_disjoint: "equivp R \ C \ UNIV // {(x, y). R x y} \ rel_set R A B \ A \ C = {} \ B \ C = {}" - using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C] + using in_quotient_imp_closed[of UNIV "{(x, y). R x y}" C] by (auto 0 0 simp: equivp_equiv rel_set_def set_eq_iff elim: equivpE) (blast dest: equivp_symp)+ @@ -973,17 +973,17 @@ next fix C assume C: "C \ UNIV // ?R" and R: "rel_set R (set_pmf p) (set_pmf q)" assume eq: "\x\set_pmf p. \y\set_pmf q. R x y \ measure p {x. R x y} = measure q {y. R x y}" - + show "measure p C = measure q C" proof cases assume "p \ C = {}" - moreover then have "q \ C = {}" + moreover then have "q \ C = {}" using quotient_rel_set_disjoint[OF assms C R] by simp ultimately show ?thesis unfolding measure_pmf_zero_iff[symmetric] by simp next assume "p \ C \ {}" - moreover then have "q \ C \ {}" + moreover then have "q \ C \ {}" using quotient_rel_set_disjoint[OF assms C R] by simp ultimately obtain x y where in_set: "x \ set_pmf p" "y \ set_pmf q" and in_C: "x \ C" "y \ C" by auto @@ -1068,11 +1068,11 @@ and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto from qr obtain qr where qr: "\y z. (y, z) \ set_pmf qr \ S y z" and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto - + def pr \ "bind_pmf pq (\xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\yz. return_pmf (fst xy, snd yz)))" have pr_welldefined: "\y. y \ q \ qr \ {yz. fst yz = y} \ {}" by (force simp: q') - + have "rel_pmf (R OO S) p r" proof (rule rel_pmf.intros) fix x z assume "(x, z) \ pr" @@ -1283,7 +1283,7 @@ proof (subst rel_pmf_iff_equivp, safe) show "equivp (inf R R\\)" using trans refl by (auto simp: equivp_reflp_symp_transp intro: sympI transpI reflpI dest: transpD reflpD) - + fix C assume "C \ UNIV // {(x, y). inf R R\\ x y}" then obtain x where C: "C = {y. R x y \ R y x}" by (auto elim: quotientE) @@ -1399,7 +1399,7 @@ end lemma set_pmf_geometric: "0 < p \ p < 1 \ set_pmf (geometric_pmf p) = UNIV" - by (auto simp: set_pmf_iff) + by (auto simp: set_pmf_iff) subsubsection \ Uniform Multiset Distribution \ @@ -1445,7 +1445,7 @@ lemma emeasure_pmf_of_set[simp]: "emeasure pmf_of_set S = 1" by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff) -lemma nn_integral_pmf_of_set': +lemma nn_integral_pmf_of_set': "(\x. x \ S \ f x \ 0) \ nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S" apply(subst nn_integral_measure_pmf_finite, simp_all add: S_finite) apply(simp add: setsum_ereal_left_distrib[symmetric]) @@ -1453,7 +1453,7 @@ apply(simp add: ereal_times_divide_eq one_ereal_def[symmetric]) done -lemma nn_integral_pmf_of_set: +lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = setsum (max 0 \ f) S / card S" apply(subst nn_integral_max_0[symmetric]) apply(subst nn_integral_pmf_of_set') @@ -1476,7 +1476,7 @@ lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x" by(rule pmf_eqI)(simp add: indicator_def) -lemma map_pmf_of_set_inj: +lemma map_pmf_of_set_inj: assumes f: "inj_on f A" and [simp]: "A \ {}" "finite A" shows "map_pmf f (pmf_of_set A) = pmf_of_set (f ` A)" (is "?lhs = ?rhs") @@ -1540,7 +1540,7 @@ ereal (\k\n. real (n choose k) * p ^ k * (1 - p) ^ (n - k))" using p_le_1 p_nonneg by (subst nn_integral_count_space') auto also have "(\k\n. real (n choose k) * p ^ k * (1 - p) ^ (n - k)) = (p + (1 - p)) ^ n" - by (subst binomial_ring) (simp add: atLeast0AtMost real_of_nat_def) + by (subst binomial_ring) (simp add: atLeast0AtMost) finally show "(\\<^sup>+ x. ereal (real (n choose x) * p ^ x * (1 - p) ^ (n - x)) \count_space UNIV) = 1" by simp qed (insert p_nonneg p_le_1, simp) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Projective_Limit.thy --- a/src/HOL/Probability/Projective_Limit.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Projective_Limit.thy Tue Nov 10 14:18:41 2015 +0000 @@ -293,9 +293,9 @@ also have "\ \ ereal (2 powr - real i) * ?a" using K'(1)[of i] . finally show "\G (Z i - Z' i) \ (2 powr - real i) * ?a" . qed - also have "\ = ereal ((\ i\{1..n}. (2 powr -real i)) * real ?a)" + also have "\ = ereal ((\ i\{1..n}. (2 powr -real_of_ereal i)) * real_of_ereal ?a)" by (simp add: setsum_left_distrib r) - also have "\ < ereal (1 * real ?a)" unfolding less_ereal.simps + also have "\ < ereal (1 * real_of_ereal ?a)" unfolding less_ereal.simps proof (rule mult_strict_right_mono) have "(\i = 1..n. 2 powr - real i) = (\i = 1.. < 1" by (subst geometric_sum) auto - finally show "(\i = 1..n. 2 powr - real i) < 1" . + finally show "(\i = 1..n. 2 powr - real_of_ereal i) < 1" by simp qed (auto simp: r ereal_less_real_iff zero_ereal_def[symmetric]) also have "\ = ?a" by (auto simp: r) also have "\ \ \G (Z n)" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Radon_Nikodym.thy Tue Nov 10 14:18:41 2015 +0000 @@ -11,7 +11,7 @@ definition "diff_measure M N = measure_of (space M) (sets M) (\A. emeasure M A - emeasure N A)" -lemma +lemma shows space_diff_measure[simp]: "space (diff_measure M N) = space M" and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M" by (auto simp: diff_measure_def) @@ -205,7 +205,7 @@ by (auto simp add: not_less) { fix n have "?d (A n) \ - real n * e" proof (induct n) - case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps) + case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: of_nat_Suc field_simps) next case 0 with measure_empty show ?case by (simp add: zero_ereal_def) qed } note dA_less = this @@ -272,7 +272,7 @@ hence "0 < - ?d B" by auto from ex_inverse_of_nat_Suc_less[OF this] obtain n where *: "?d B < - 1 / real (Suc n)" - by (auto simp: real_eq_of_nat field_simps) + by (auto simp: field_simps) also have "\ \ - 1 / real (Suc (Suc n))" by (auto simp: field_simps) finally show False @@ -443,8 +443,8 @@ using b by (auto simp: emeasure_density_const sets_eq intro!: finite_measureI) from M'.Radon_Nikodym_aux[OF this] guess A0 .. then have "A0 \ sets M" - and space_less_A0: "measure ?M (space M) - real b * measure M (space M) \ measure ?M A0 - real b * measure M A0" - and *: "\B. B \ sets M \ B \ A0 \ 0 \ measure ?M B - real b * measure M B" + and space_less_A0: "measure ?M (space M) - real_of_ereal b * measure M (space M) \ measure ?M A0 - real_of_ereal b * measure M A0" + and *: "\B. B \ sets M \ B \ A0 \ 0 \ measure ?M B - real_of_ereal b * measure M B" using b by (simp_all add: measure_density_const sets_eq_imp_space_eq[OF sets_eq] sets_eq) { fix B assume B: "B \ sets M" "B \ A0" with *[OF this] have "b * emeasure M B \ ?M B" @@ -820,7 +820,7 @@ let ?f = "\A x. f x * indicator A x" and ?f' = "\A x. f' x * indicator A x" have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M" - using borel by (auto intro!: absolutely_continuousI_density) + using borel by (auto intro!: absolutely_continuousI_density) from split_space_into_finite_sets_and_rest[OF this] obtain Q0 and Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" @@ -857,7 +857,7 @@ then show "?A i \ null_sets M" using in_Q0[OF `?A i \ sets M`] `?A i \ sets M` by auto qed also have "(\i. ?A i) = Q0 \ {x\space M. f x \ \}" - by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat) + by (auto simp: less_PInf_Ex_of_nat) finally have "Q0 \ {x\space M. f x \ \} \ null_sets M" by simp } from this[OF borel(1) refl] this[OF borel(2) f] have "Q0 \ {x\space M. f x \ \} \ null_sets M" "Q0 \ {x\space M. f' x \ \} \ null_sets M" by simp_all @@ -931,7 +931,7 @@ and fin: "sigma_finite_measure (density M f)" shows "density M f = density M g \ (AE x in M. f x = g x)" proof - assume "AE x in M. f x = g x" with borel show "density M f = density M g" + assume "AE x in M. f x = g x" with borel show "density M f = density M g" by (auto intro: density_cong) next assume eq: "density M f = density M g" @@ -1010,8 +1010,8 @@ case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0]) next case (real r) - with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by (auto simp: real_eq_of_nat) - then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"] simp: real_eq_of_nat) + with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by auto + then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"]) next case MInf with x show ?thesis unfolding A_def by (auto intro!: exI[of _ "Suc 0"]) @@ -1031,11 +1031,11 @@ case (Suc n) then have "(\\<^sup>+x. f x * indicator (A i \ Q j) x \M) \ (\\<^sup>+x. (Suc n :: ereal) * indicator (Q j) x \M)" - by (auto intro!: nn_integral_mono simp: indicator_def A_def real_eq_of_nat) + by (auto intro!: nn_integral_mono simp: indicator_def A_def) also have "\ = Suc n * emeasure M (Q j)" using Q by (auto intro!: nn_integral_cmult_indicator) also have "\ < \" - using Q by (auto simp: real_eq_of_nat[symmetric]) + using Q by auto finally show ?thesis by simp qed then show "emeasure ?N X \ \" @@ -1058,7 +1058,7 @@ then SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N else (\_. 0))" -lemma RN_derivI: +lemma RN_derivI: assumes "f \ borel_measurable M" "\x. 0 \ f x" "density M f = N" shows "density M (RN_deriv M N) = N" proof - @@ -1070,7 +1070,7 @@ by (auto simp: RN_deriv_def) qed -lemma +lemma shows borel_measurable_RN_deriv[measurable]: "RN_deriv M N \ borel_measurable M" (is ?m) and RN_deriv_nonneg: "0 \ RN_deriv M N x" (is ?nn) proof - @@ -1205,16 +1205,16 @@ assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" and f: "f \ borel_measurable M" shows RN_deriv_integrable: "integrable N f \ - integrable M (\x. real (RN_deriv M N x) * f x)" (is ?integrable) - and RN_deriv_integral: "integral\<^sup>L N f = (\x. real (RN_deriv M N x) * f x \M)" (is ?integral) + integrable M (\x. real_of_ereal (RN_deriv M N x) * f x)" (is ?integrable) + and RN_deriv_integral: "integral\<^sup>L N f = (\x. real_of_ereal (RN_deriv M N x) * f x \M)" (is ?integral) proof - note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp] interpret N: sigma_finite_measure N by fact - have eq: "density M (RN_deriv M N) = density M (\x. real (RN_deriv M N x))" + have eq: "density M (RN_deriv M N) = density M (\x. real_of_ereal (RN_deriv M N x))" proof (rule density_cong) from RN_deriv_finite[OF assms(1,2,3)] - show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))" + show "AE x in M. RN_deriv M N x = ereal (real_of_ereal (RN_deriv M N x))" by eventually_elim (insert RN_deriv_nonneg[of M N], auto simp: ereal_real) qed (insert ac, auto) @@ -1242,12 +1242,12 @@ and "\x. 0 \ D x" proof interpret N: finite_measure N by fact - + note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac] RN_deriv_nonneg[of M N] let ?RN = "\t. {x \ space M. RN_deriv M N x = t}" - show "(\x. real (RN_deriv M N x)) \ borel_measurable M" + show "(\x. real_of_ereal (RN_deriv M N x)) \ borel_measurable M" using RN by auto have "N (?RN \) = (\\<^sup>+ x. RN_deriv M N x * indicator (?RN \) x \M)" @@ -1268,12 +1268,12 @@ qed ultimately have "AE x in M. RN_deriv M N x < \" using RN by (intro AE_iff_measurable[THEN iffD2]) auto - then show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))" + then show "AE x in M. RN_deriv M N x = ereal (real_of_ereal (RN_deriv M N x))" using RN(3) by (auto simp: ereal_real) - then have eq: "AE x in N. RN_deriv M N x = ereal (real (RN_deriv M N x))" + then have eq: "AE x in N. RN_deriv M N x = ereal (real_of_ereal (RN_deriv M N x))" using ac absolutely_continuous_AE by auto - show "\x. 0 \ real (RN_deriv M N x)" + show "\x. 0 \ real_of_ereal (RN_deriv M N x)" using RN by (auto intro: real_of_ereal_pos) have "N (?RN 0) = (\\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \M)" @@ -1282,7 +1282,7 @@ by (intro nn_integral_cong) (auto simp: indicator_def) finally have "AE x in N. RN_deriv M N x \ 0" using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq) - with RN(3) eq show "AE x in N. 0 < real (RN_deriv M N x)" + with RN(3) eq show "AE x in N. 0 < real_of_ereal (RN_deriv M N x)" by (auto simp: zero_less_real_of_ereal le_less) qed diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Regularity.thy --- a/src/HOL/Probability/Regularity.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Regularity.thy Tue Nov 10 14:18:41 2015 +0000 @@ -26,7 +26,7 @@ assume "\ x \ y" hence "x > y" by simp hence y_fin: "\y\ \ \" using `y \ 0` by auto have x_fin: "\x\ \ \" using `x > y` f_fin approx[where e = 1] by auto - def e \ "real ((x - y) / 2)" + def e \ "real_of_ereal ((x - y) / 2)" have e: "x > y + e" "e > 0" using `x > y` y_fin x_fin by (auto simp: e_def field_simps) note e(1) also from approx[OF `e > 0`] obtain i where i: "i \ A" "x \ f i + e" by blast @@ -56,7 +56,7 @@ hence y_fin: "\y\ \ \" using `y \ \` by auto have x_fin: "\x\ \ \" using `y > x` f_fin f_nonneg approx[where e = 1] A_notempty by auto - def e \ "real ((y - x) / 2)" + def e \ "real_of_ereal ((y - x) / 2)" have e: "y > x + e" "e > 0" using `y > x` y_fin x_fin by (auto simp: e_def field_simps) from approx[OF `e > 0`] obtain i where i: "i \ A" "x + e \ f i" by blast note i(2) @@ -192,7 +192,7 @@ also have "\ \ (\n. ereal (e*2 powr - real (Suc n)))" using B_compl_le by (intro suminf_le_pos) (simp_all add: measure_nonneg emeasure_eq_measure) also have "\ \ (\n. ereal (e * (1 / 2) ^ Suc n))" - by (simp add: Transcendental.powr_minus powr_realpow field_simps del: real_of_nat_Suc) + by (simp add: Transcendental.powr_minus powr_realpow field_simps del: of_nat_Suc) also have "\ = (\n. ereal e * ((1 / 2) ^ Suc n))" unfolding times_ereal.simps[symmetric] ereal_power[symmetric] one_ereal_def numeral_eq_ereal by simp @@ -254,7 +254,7 @@ from in_closed_iff_infdist_zero[OF `closed A` `A \ {}`] have "A = {x. infdist x A = 0}" by auto also have "\ = (\i. ?G (1/real (Suc i)))" - proof (auto simp del: real_of_nat_Suc, rule ccontr) + proof (auto simp del: of_nat_Suc, rule ccontr) fix x assume "infdist x A \ 0" hence pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp @@ -262,7 +262,7 @@ moreover assume "\i. infdist x A < 1 / real (Suc i)" hence "infdist x A < 1 / real (Suc n)" by auto - ultimately show False by simp + ultimately show False by simp qed also have "M \ = (INF n. emeasure M (?G (1 / real (Suc n))))" proof (rule INF_emeasure_decseq[symmetric], safe) @@ -294,7 +294,7 @@ note `?inner A` `?outer A` } note closed_in_D = this from `B \ sets borel` - have "Int_stable (Collect closed)" "Collect closed \ Pow UNIV" "B \ sigma_sets UNIV (Collect closed)" + have "Int_stable (Collect closed)" "Collect closed \ Pow UNIV" "B \ sigma_sets UNIV (Collect closed)" by (auto simp: Int_stable_def borel_eq_closed) then show "?inner B" "?outer B" proof (induct B rule: sigma_sets_induct_disjoint) @@ -356,7 +356,7 @@ finally have measure_LIMSEQ: "(\n. \i measure M (\i. D i)" by (simp add: emeasure_eq_measure) have "(\i. D i) \ sets M" using `range D \ sets M` by auto - + case 1 show ?case proof (rule approx_inner) @@ -398,7 +398,7 @@ also have "\ = (\ii \ (\ii. D i) < (\i \ (\n. ereal (e * (1 / 2) ^ Suc n))" - by (simp add: powr_minus inverse_eq_divide powr_realpow field_simps power_divide del: real_of_nat_Suc) + by (simp add: powr_minus inverse_eq_divide powr_realpow field_simps power_divide del: of_nat_Suc) also have "\ = (\n. ereal e * ((1 / 2) ^ Suc n))" unfolding times_ereal.simps[symmetric] ereal_power[symmetric] one_ereal_def numeral_eq_ereal by simp diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Tue Nov 10 14:18:41 2015 +0000 @@ -1468,7 +1468,7 @@ "emeasure M = snd (snd (Rep_measure M))" definition measure :: "'a measure \ 'a set \ real" where - "measure M A = real (emeasure M A)" + "measure M A = real_of_ereal (emeasure M A)" declare [[coercion sets]] diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Tue Nov 10 14:18:41 2015 +0000 @@ -374,7 +374,7 @@ using finite_measure_finite_Union[OF _ df] by (auto simp add: * intro!: setsum_nonneg) also have "(\obs'\?S obs. \

(OB ; fst) {(obs', k)}) = real (card (?S obs)) * \

(OB ; fst) {(obs, k)}" - by (simp add: t_eq_imp[OF `k \ keys` `K k \ 0` obs] real_eq_of_nat) + by (simp add: t_eq_imp[OF `k \ keys` `K k \ 0` obs]) finally have "\

(t\OB ; fst) {(t obs, k)} = real (card (?S obs)) * \

(OB ; fst) {(obs, k)}" .} note P_t_eq_P_OB = this @@ -471,9 +471,9 @@ using entropy_le_card[of "t\OB", OF simple_distributedI[OF simple_function_finite refl]] by simp also have "\ \ log b (real (n + 1)^card observations)" using card_T_bound not_empty - by (auto intro!: log_le simp: card_gt_0_iff power_real_of_nat simp del: real_of_nat_power real_of_nat_Suc) + by (auto intro!: log_le simp: card_gt_0_iff power_real_of_nat simp del: of_nat_power of_nat_Suc) also have "\ = real (card observations) * log b (real n + 1)" - by (simp add: log_nat_power real_of_nat_Suc) + by (simp add: log_nat_power add.commute) finally show ?thesis . qed diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Real.thy --- a/src/HOL/Real.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Real.thy Tue Nov 10 14:18:41 2015 +0000 @@ -760,10 +760,9 @@ shows "x < Real Y \ \n. x < of_rat (Y n)" by (erule contrapos_pp, simp add: not_less, erule Real_leI [OF Y]) -lemma of_nat_less_two_power: +lemma of_nat_less_two_power [simp]: "of_nat n < (2::'a::linordered_idom) ^ n" -apply (induct n) -apply simp +apply (induct n, simp) by (metis add_le_less_mono mult_2 of_nat_Suc one_le_numeral one_le_power power_Suc) lemma complete_real: @@ -995,6 +994,11 @@ "real_of_nat \ of_nat" abbreviation + real :: "nat \ real" +where + "real \ of_nat" + +abbreviation real_of_int :: "int \ real" where "real_of_int \ of_int" @@ -1004,30 +1008,11 @@ where "real_of_rat \ of_rat" -class real_of = - fixes real :: "'a \ real" - -instantiation nat :: real_of -begin - -definition real_nat :: "nat \ real" where real_of_nat_def [code_unfold]: "real \ of_nat" - -instance .. -end - -instantiation int :: real_of -begin - -definition real_int :: "int \ real" where real_of_int_def [code_unfold]: "real \ of_int" - -instance .. -end - declare [[coercion_enabled]] declare [[coercion "of_nat :: nat \ int"]] -declare [[coercion "real :: nat \ real"]] -declare [[coercion "real :: int \ real"]] +declare [[coercion "of_nat :: nat \ real"]] +declare [[coercion "of_int :: int \ real"]] (* We do not add rat to the coerced types, this has often unpleasant side effects when writing inverse (Suc n) which sometimes gets two coercions: of_rat (inverse (of_nat (Suc n))) *) @@ -1036,105 +1021,40 @@ declare [[coercion_map "\f g h x. g (h (f x))"]] declare [[coercion_map "\f g (x,y). (f x, g y)"]] -lemma real_eq_of_nat: "real = of_nat" - unfolding real_of_nat_def .. - -lemma real_eq_of_int: "real = of_int" - unfolding real_of_int_def .. - -lemma real_of_int_zero [simp]: "real (0::int) = 0" -by (simp add: real_of_int_def) - -lemma real_of_one [simp]: "real (1::int) = (1::real)" -by (simp add: real_of_int_def) - -lemma real_of_int_add [simp]: "real(x + y) = real (x::int) + real y" -by (simp add: real_of_int_def) - -lemma real_of_int_minus [simp]: "real(-x) = -real (x::int)" -by (simp add: real_of_int_def) - -lemma real_of_int_diff [simp]: "real(x - y) = real (x::int) - real y" -by (simp add: real_of_int_def) +declare of_int_eq_0_iff [algebra, presburger] +declare of_int_eq_iff [iff, algebra, presburger] (*FIXME*) +declare of_int_eq_iff [iff, algebra, presburger] (*FIXME*) +declare of_int_less_iff [iff, algebra, presburger] (*FIXME*) +declare of_int_le_iff [iff, algebra, presburger] (*FIXME*) -lemma real_of_int_mult [simp]: "real(x * y) = real (x::int) * real y" -by (simp add: real_of_int_def) - -lemma real_of_int_power [simp]: "real (x ^ n) = real (x::int) ^ n" -by (simp add: real_of_int_def of_int_power) - -lemmas power_real_of_int = real_of_int_power [symmetric] +declare of_int_0_less_iff [presburger] +declare of_int_0_le_iff [presburger] +declare of_int_less_0_iff [presburger] +declare of_int_le_0_iff [presburger] -lemma real_of_int_setsum [simp]: "real ((SUM x:A. f x)::int) = (SUM x:A. real(f x))" - apply (subst real_eq_of_int)+ - apply (rule of_int_setsum) -done - -lemma real_of_int_setprod [simp]: "real ((PROD x:A. f x)::int) = - (PROD x:A. real(f x))" - apply (subst real_eq_of_int)+ - apply (rule of_int_setprod) -done - -lemma real_of_int_zero_cancel [simp, algebra, presburger]: "(real x = 0) = (x = (0::int))" -by (simp add: real_of_int_def) - -lemma real_of_int_inject [iff, algebra, presburger]: "(real (x::int) = real y) = (x = y)" -by (simp add: real_of_int_def) +lemma real_of_int_abs [simp]: "real_of_int (abs x) = abs(real_of_int x)" + by (auto simp add: abs_if) -lemma real_of_int_less_iff [iff, presburger]: "(real (x::int) < real y) = (x < y)" -by (simp add: real_of_int_def) - -lemma real_of_int_le_iff [simp, presburger]: "(real (x::int) \ real y) = (x \ y)" -by (simp add: real_of_int_def) - -lemma real_of_int_gt_zero_cancel_iff [simp, presburger]: "(0 < real (n::int)) = (0 < n)" -by (simp add: real_of_int_def) - -lemma real_of_int_ge_zero_cancel_iff [simp, presburger]: "(0 <= real (n::int)) = (0 <= n)" -by (simp add: real_of_int_def) - -lemma real_of_int_lt_zero_cancel_iff [simp, presburger]: "(real (n::int) < 0) = (n < 0)" -by (simp add: real_of_int_def) - -lemma real_of_int_le_zero_cancel_iff [simp, presburger]: "(real (n::int) <= 0) = (n <= 0)" -by (simp add: real_of_int_def) - -lemma one_less_real_of_int_cancel_iff: "1 < real (i :: int) \ 1 < i" - unfolding real_of_one[symmetric] real_of_int_less_iff .. - -lemma one_le_real_of_int_cancel_iff: "1 \ real (i :: int) \ 1 \ i" - unfolding real_of_one[symmetric] real_of_int_le_iff .. +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 + by (metis floor_unique less_add_one less_imp_le not_less of_int_le_iff order_trans) +qed -lemma real_of_int_less_one_cancel_iff: "real (i :: int) < 1 \ i < 1" - unfolding real_of_one[symmetric] real_of_int_less_iff .. - -lemma real_of_int_le_one_cancel_iff: "real (i :: int) \ 1 \ i \ 1" - unfolding real_of_one[symmetric] real_of_int_le_iff .. - -lemma real_of_int_abs [simp]: "real (abs x) = abs(real (x::int))" -by (auto simp add: abs_if) +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 int_less_real_le: "((n::int) < m) = (real n + 1 <= real m)" - apply (subgoal_tac "real n + 1 = real (n + 1)") - apply (simp del: real_of_int_add) - apply auto -done -lemma int_le_real_less: "((n::int) <= m) = (real n < real m + 1)" - apply (subgoal_tac "real m + 1 = real (m + 1)") - apply (simp del: real_of_int_add) - apply simp -done - -lemma real_of_int_div_aux: "(real (x::int)) / (real d) = - real (x div d) + (real (x mod d)) / (real 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 x = real (x div d) * real d + real(x mod d)" - by (simp only: real_of_int_mult [THEN sym] real_of_int_add [THEN sym]) - then have "real x / real d = ... / real d" + 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" by simp then show ?thesis by (auto simp add: add_divide_distrib algebra_simps) @@ -1142,133 +1062,73 @@ lemma real_of_int_div: fixes d n :: int - shows "d dvd n \ real (n div d) = real n / real d" + shows "d dvd n \ real_of_int (n div d) = real_of_int n / real_of_int d" by (simp add: real_of_int_div_aux) lemma real_of_int_div2: - "0 <= real (n::int) / real (x) - real (n div x)" - apply (case_tac "x = 0") - apply simp + "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") - apply (simp add: algebra_simps) - apply (subst real_of_int_div_aux) - apply simp - apply (simp add: algebra_simps) - apply (subst real_of_int_div_aux) - apply simp - apply (subst zero_le_divide_iff) - apply auto -done + 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 (n::int) / real (x) - real (n div x) <= 1" + "real_of_int (n::int) / 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 -lemma real_of_int_div4: "real (n div x) <= real (n::int) / real x" +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 Ints_real_of_int [simp]: "real (x::int) \ \" -unfolding real_of_int_def by (rule Ints_of_int) - subsection\Embedding the Naturals into the Reals\ -lemma real_of_nat_zero [simp]: "real (0::nat) = 0" -by (simp add: real_of_nat_def) - -lemma real_of_nat_1 [simp]: "real (1::nat) = 1" -by (simp add: real_of_nat_def) - -lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_add [simp]: "real (m + n) = real (m::nat) + real n" -by (simp add: real_of_nat_def) - -lemma real_of_nat_Suc [simp]: "real (Suc n) = real n + (1::real)" -by (simp add: real_of_nat_def) +declare of_nat_less_iff [iff] (*FIXME*) +declare of_nat_le_iff [iff] (*FIXME*) +declare of_nat_0_le_iff [iff] (*FIXME*) +declare of_nat_less_iff [iff] (*FIXME*) +declare of_nat_less_iff [iff] (*FIXME*) -lemma real_of_nat_less_iff [iff]: - "(real (n::nat) < real m) = (n < m)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_le_iff [iff]: "(real (n::nat) \ real m) = (n \ m)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_ge_zero [iff]: "0 \ real (n::nat)" -by (simp add: real_of_nat_def) +lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)" (*NEEDED?*) + using of_nat_0_less_iff by blast -lemma real_of_nat_Suc_gt_zero: "0 < real (Suc n)" -by (simp add: real_of_nat_def del: of_nat_Suc) - -lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n" -by (simp add: real_of_nat_def of_nat_mult) - -lemma real_of_nat_power [simp]: "real (m ^ n) = real (m::nat) ^ n" -by (simp add: real_of_nat_def of_nat_power) - -lemmas power_real_of_nat = real_of_nat_power [symmetric] +declare of_nat_mult [simp] (*FIXME*) +declare of_nat_power [simp] (*FIXME*) -lemma real_of_nat_setsum [simp]: "real ((SUM x:A. f x)::nat) = - (SUM x:A. real(f x))" - apply (subst real_eq_of_nat)+ - apply (rule of_nat_setsum) -done +lemmas power_real_of_nat = of_nat_power [symmetric] -lemma real_of_nat_setprod [simp]: "real ((PROD x:A. f x)::nat) = - (PROD x:A. real(f x))" - apply (subst real_eq_of_nat)+ - apply (rule of_nat_setprod) -done - -lemma real_of_card: "real (card A) = setsum (%x.1) A" - apply (subst card_eq_setsum) - apply (subst real_of_nat_setsum) - apply simp -done +declare of_nat_setsum [simp] (*FIXME*) +declare of_nat_setprod [simp] (*FIXME*) -lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_zero_iff [iff]: "(real (n::nat) = 0) = (n = 0)" -by (simp add: real_of_nat_def) - -lemma real_of_nat_diff: "n \ m ==> real (m - n) = real (m::nat) - real n" -by (simp add: add: real_of_nat_def of_nat_diff) +lemma real_of_card: "real (card A) = setsum (\x.1) A" + by simp -lemma real_of_nat_gt_zero_cancel_iff [simp]: "(0 < real (n::nat)) = (0 < n)" -by (auto simp: real_of_nat_def) - -lemma real_of_nat_le_zero_cancel_iff [simp]: "(real (n::nat) \ 0) = (n = 0)" -by (simp add: add: real_of_nat_def) +declare of_nat_eq_iff [iff] (*FIXME*) +declare of_nat_eq_0_iff [iff] (*FIXME*) -lemma not_real_of_nat_less_zero [simp]: "~ real (n::nat) < 0" -by (simp add: add: real_of_nat_def) - -lemma nat_less_real_le: "((n::nat) < m) = (real n + 1 <= real m)" -by (metis discrete real_of_nat_1 real_of_nat_add real_of_nat_le_iff) +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)" by (meson nat_less_real_le not_le) -lemma real_of_nat_div_aux: "(real (x::nat)) / (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 then have "real x = real (x div d) * real d + real(x mod d)" - by (simp only: real_of_nat_mult [THEN sym] real_of_nat_add [THEN sym]) + by (metis of_nat_add of_nat_mult) then have "real x / real d = \ / real d" by simp then show ?thesis by (auto simp add: add_divide_distrib algebra_simps) qed -lemma real_of_nat_div: "(d :: nat) dvd n ==> - real(n div d) = real n / real d" +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]) @@ -1291,85 +1151,63 @@ 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_int_of_nat_eq [simp]: "real (of_nat n :: int) = real n" -by (simp add: real_of_int_def real_of_nat_def) - -lemma real_nat_eq_real [simp]: "0 <= x ==> real(nat x) = real x" - apply (subgoal_tac "real(int(nat x)) = real(nat x)") - apply force - apply (simp only: real_of_int_of_nat_eq) -done - -lemma Nats_real_of_nat [simp]: "real (n::nat) \ \" -unfolding real_of_nat_def by (rule of_nat_in_Nats) - -lemma Ints_real_of_nat [simp]: "real (n::nat) \ \" -unfolding real_of_nat_def by (rule Ints_of_nat) +lemma of_nat_nat [simp]: "0 <= x ==> real(nat x) = real x" + by force subsection \The Archimedean Property of the Reals\ -theorem reals_Archimedean: - assumes x_pos: "0 < x" - shows "\n. inverse (real (Suc n)) < x" - unfolding real_of_nat_def using x_pos - by (rule ex_inverse_of_nat_Suc_less) - -lemma reals_Archimedean2: "\n. (x::real) < real (n::nat)" - unfolding real_of_nat_def by (rule ex_less_of_nat) +lemmas reals_Archimedean = ex_inverse_of_nat_Suc_less (*FIXME*) +lemmas reals_Archimedean2 = ex_less_of_nat lemma reals_Archimedean3: assumes x_greater_zero: "0 < x" - shows "\(y::real). \(n::nat). y < real n * x" - unfolding real_of_nat_def using \0 < x\ - by (auto intro: ex_less_of_nat_mult) + shows "\y. \n. y < real n * x" + using \0 < x\ by (auto intro: ex_less_of_nat_mult) subsection\Rationals\ -lemma Rats_real_nat[simp]: "real(n::nat) \ \" -by (simp add: real_eq_of_nat) - lemma Rats_eq_int_div_int: - "\ = { real(i::int)/real(j::int) |i j. j \ 0}" (is "_ = ?S") + "\ = { 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 real_eq_of_int) + by (cases r) (auto simp add:of_rat_rat) thus "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 i / real j = of_rat(Fract i j)" - by (simp add:of_rat_rat real_eq_of_int) - thus "real i / real j \ range of_rat" by blast + hence "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 qed qed lemma Rats_eq_int_div_nat: - "\ = { real(i::int)/real(n::nat) |i n. n \ 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 "EX (i'::int) (n::nat). real i/real j = real i'/real n \ 0 00" - hence "real i/real j = real i/real(nat j) \ 0 00" - hence "real i/real j = real(-i)/real(nat(-j)) \ 0j\0\ - by (simp add: real_eq_of_int real_eq_of_nat of_nat_nat) + 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 qed next fix i::int and n::nat assume "0 < n" - hence "real i/real n = real i/real(int n) \ int n \ 0" by simp - thus "\(i'::int) j::int. real i/real n = real i'/real j \ j \ 0" by blast + 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 qed lemma Rats_abs_nat_div_natE: @@ -1377,9 +1215,9 @@ 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 i / real n" + 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(abs i)) / real n" by simp + hence "\x\ = real(nat(abs i)) / real n" by (simp add: of_nat_nat) 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 @@ -1396,9 +1234,8 @@ moreover have "\x\ = real ?k / real ?l" proof - - from gcd have "real ?k / real ?l = - real (?gcd * ?k) / real (?gcd * ?l)" - by (simp only: real_of_nat_mult) simp + from gcd have "real ?k / real ?l = real (?gcd * ?k) / real (?gcd * ?l)" + by (simp add: real_of_nat_div) also from gcd_k and gcd_l have "\ = real m / real n" by simp also from x_rat have "\ = \x\" .. finally show ?thesis .. @@ -1426,7 +1263,7 @@ 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 + where q: "inverse (real q) < y-x" and "0 < q" by blast def p \ "ceiling (y * real q) - 1" def r \ "of_int p / real q" from q have "x < y - inverse (real q)" by simp @@ -1452,34 +1289,23 @@ subsection\Numerals and Arithmetic\ -lemma [code_abbrev]: +lemma [code_abbrev]: (*FIXME*) "real_of_int (numeral k) = numeral k" "real_of_int (- numeral k) = - numeral k" by simp_all -text\Collapse applications of @{const real} to @{const numeral}\ -lemma real_numeral [simp]: - "real (numeral v :: int) = numeral v" - "real (- numeral v :: int) = - numeral v" -by (simp_all add: real_of_int_def) -lemma real_of_nat_numeral [simp]: - "real (numeral v :: nat) = numeral v" -by (simp add: real_of_nat_def) - + (*FIXME*) declaration \ - K (Lin_Arith.add_inj_thms [@{thm real_of_nat_le_iff} RS iffD2, @{thm real_of_nat_inject} RS iffD2] - (* not needed because x < (y::nat) can be rewritten as Suc x <= y: real_of_nat_less_iff RS iffD2 *) - #> Lin_Arith.add_inj_thms [@{thm real_of_int_le_iff} RS iffD2, @{thm real_of_int_inject} RS iffD2] - (* not needed because x < (y::int) can be rewritten as x + 1 <= y: real_of_int_less_iff RS iffD2 *) - #> Lin_Arith.add_simps [@{thm real_of_nat_zero}, @{thm real_of_nat_Suc}, @{thm real_of_nat_add}, - @{thm real_of_nat_mult}, @{thm real_of_int_zero}, @{thm real_of_one}, - @{thm real_of_int_add}, @{thm real_of_int_minus}, @{thm real_of_int_diff}, - @{thm real_of_int_mult}, @{thm real_of_int_of_nat_eq}, - @{thm real_of_nat_numeral}, @{thm real_numeral(1)}, @{thm real_numeral(2)}, - @{thm real_of_int_def[symmetric]}, @{thm real_of_nat_def[symmetric]}] - #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "nat \ real"}) - #> Lin_Arith.add_inj_const (@{const_name real}, @{typ "int \ real"}) + K (Lin_Arith.add_inj_thms [@{thm of_nat_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2] + (* not needed because x < (y::nat) can be rewritten as Suc x <= y: of_nat_less_iff RS iffD2 *) + #> Lin_Arith.add_inj_thms [@{thm of_int_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2] + (* not needed because x < (y::int) can be rewritten as x + 1 <= y: of_int_less_iff RS iffD2 *) + #> Lin_Arith.add_simps [@{thm of_nat_0}, @{thm of_nat_Suc}, @{thm of_nat_add}, + @{thm of_nat_mult}, @{thm of_int_0}, @{thm of_int_1}, + @{thm of_int_add}, @{thm of_int_minus}, @{thm of_int_diff}, + @{thm of_int_mult}, @{thm of_int_of_nat_eq}, + @{thm of_nat_numeral}, @{thm int_numeral}, @{thm of_int_neg_numeral}] #> Lin_Arith.add_inj_const (@{const_name of_nat}, @{typ "nat \ real"}) #> Lin_Arith.add_inj_const (@{const_name of_int}, @{typ "int \ real"})) \ @@ -1489,10 +1315,6 @@ lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" by arith -text \FIXME: redundant with @{text add_eq_0_iff} below\ -lemma real_add_eq_0_iff: "(x+y = (0::real)) = (y = -x)" -by auto - lemma real_add_less_0_iff: "(x+y < (0::real)) = (y < -x)" by auto @@ -1512,53 +1334,41 @@ (* used by Import/HOL/real.imp *) lemma two_realpow_ge_one: "(1::real) \ 2 ^ n" -by simp + by simp -lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n" - by (simp add: of_nat_less_two_power real_of_nat_def) - -text \TODO: no longer real-specific; rename and move elsewhere\ +text \FIXME: no longer real-specific; rename and move elsewhere\ lemma realpow_Suc_le_self: fixes r :: "'a::linordered_semidom" shows "[| 0 \ r; r \ 1 |] ==> r ^ Suc n \ r" by (insert power_decreasing [of 1 "Suc n" r], simp) -text \TODO: no longer real-specific; rename and move elsewhere\ +text \FIXME: no longer real-specific; rename and move elsewhere\ lemma realpow_minus_mult: fixes x :: "'a::monoid_mult" shows "0 < n \ x ^ (n - 1) * x = x ^ n" by (simp add: power_Suc power_commutes split add: nat_diff_split) text \FIXME: declare this [simp] for all types, or not at all\ -lemma real_two_squares_add_zero_iff [simp]: - "(x * x + y * y = 0) = ((x::real) = 0 \ y = 0)" -by (rule sum_squares_eq_zero_iff) - -text \FIXME: declare this [simp] for all types, or not at all\ -lemma realpow_two_sum_zero_iff [simp]: - "(x\<^sup>2 + y\<^sup>2 = (0::real)) = (x = 0 & y = 0)" -by (rule sum_power2_eq_zero_iff) +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 realpow_square_minus_le [simp]: "- u\<^sup>2 \ (x::real)\<^sup>2" -by (auto simp add: power2_eq_square) - + by (auto simp add: power2_eq_square) lemma numeral_power_eq_real_of_int_cancel_iff[simp]: - "numeral x ^ n = real (y::int) \ numeral x ^ n = y" - by (metis real_numeral(1) real_of_int_inject real_of_int_power) + "numeral x ^ n = real_of_int (y::int) \ 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 (y::int) = numeral x ^ n \ y = numeral x ^ n" + "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 numeral_power_eq_real_of_nat_cancel_iff[simp]: - "numeral x ^ n = real (y::nat) \ numeral x ^ n = y" - by (metis of_nat_eq_iff of_nat_numeral real_of_int_eq_numeral_power_cancel_iff - real_of_int_of_nat_eq zpower_int) + "numeral x ^ n = real (y::nat) \ 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" @@ -1567,43 +1377,43 @@ lemma numeral_power_le_real_of_nat_cancel_iff[simp]: "(numeral x::real) ^ n \ real a \ (numeral x::nat) ^ n \ a" - unfolding real_of_nat_le_iff[symmetric] by simp +by (metis of_nat_le_iff of_nat_numeral of_nat_power) lemma real_of_nat_le_numeral_power_cancel_iff[simp]: "real a \ (numeral x::real) ^ n \ a \ (numeral x::nat) ^ n" - unfolding real_of_nat_le_iff[symmetric] by simp +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 a \ (numeral x::int) ^ n \ a" - unfolding real_of_int_le_iff[symmetric] by 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 a \ (numeral x::real) ^ n \ a \ (numeral x::int) ^ n" - unfolding real_of_int_le_iff[symmetric] by 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" - unfolding real_of_nat_less_iff[symmetric] by 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" - unfolding real_of_nat_less_iff[symmetric] by simp +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 a \ (numeral x::int) ^ n < a" - unfolding real_of_int_less_iff[symmetric] by 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 a < (numeral x::real) ^ n \ a < (numeral x::int) ^ n" - unfolding real_of_int_less_iff[symmetric] by 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 a \ (- numeral x::int) ^ n \ a" - unfolding real_of_int_le_iff[symmetric] by 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 a \ (- numeral x::real) ^ n \ a \ (- numeral x::int) ^ n" - unfolding real_of_int_le_iff[symmetric] by 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\ @@ -1637,9 +1447,6 @@ lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)" by (simp add: abs_if) -lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)" -by (rule abs_of_nonneg [OF real_of_nat_ge_zero]) - lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x" by simp @@ -1649,73 +1456,30 @@ subsection\Floor and Ceiling Functions from the Reals to the Integers\ -(* FIXME: theorems for negative numerals *) -lemma numeral_less_real_of_int_iff [simp]: - "((numeral n) < real (m::int)) = (numeral n < m)" -apply auto -apply (rule real_of_int_less_iff [THEN iffD1]) -apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) -done - -lemma numeral_less_real_of_int_iff2 [simp]: - "(real (m::int) < (numeral n)) = (m < numeral n)" -apply auto -apply (rule real_of_int_less_iff [THEN iffD1]) -apply (drule_tac [2] real_of_int_less_iff [THEN iffD2], auto) -done +(* 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" - using real_of_nat_less_iff[of n "numeral w"] by simp + "real (n::nat) < numeral w \ n < numeral w" + 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" - using real_of_nat_less_iff[of "numeral w" n] by simp + "numeral w < real (n::nat) \ numeral w < n" + 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_int_iff [simp]: - "((numeral n) \ real (m::int)) = (numeral n \ m)" -by (simp add: linorder_not_less [symmetric]) - -lemma numeral_le_real_of_int_iff2 [simp]: - "(real (m::int) \ (numeral n)) = (m \ numeral n)" -by (simp add: linorder_not_less [symmetric]) - -lemma floor_real_of_nat [simp]: "floor (real (n::nat)) = int n" -unfolding real_of_nat_def by simp - -lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n" -unfolding real_of_nat_def by (simp add: floor_minus) - -lemma floor_real_of_int [simp]: "floor (real (n::int)) = n" -unfolding real_of_int_def by simp +declare of_int_floor_le [simp] (* FIXME*) -lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n" -unfolding real_of_int_def by (simp add: floor_minus) - -lemma real_lb_ub_int: " \n::int. real n \ r & r < real (n+1)" -unfolding real_of_int_def by (rule floor_exists) - -lemma lemma_floor: "real m \ r \ r < real n + 1 \ m \ (n::int)" - by simp +lemma of_int_floor_cancel [simp]: + "(of_int (floor x) = x) = (\n::int. x = of_int n)" + by (metis floor_of_int) -lemma real_of_int_floor_le [simp]: "real (floor r) \ r" -unfolding real_of_int_def by (rule of_int_floor_le) - -lemma lemma_floor2: "real n < real (x::int) + 1 ==> n \ x" - by simp - -lemma real_of_int_floor_cancel [simp]: - "(real (floor x) = x) = (\n::int. x = real n)" - using floor_real_of_int by metis - -lemma floor_eq: "[| real n < x; x < real n + 1 |] ==> floor x = n" +lemma floor_eq: "[| real_of_int n < x; x < real_of_int n + 1 |] ==> floor x = n" by linarith -lemma floor_eq2: "[| real n \ x; x < real n + 1 |] ==> floor x = n" +lemma floor_eq2: "[| real_of_int n \ x; x < real_of_int n + 1 |] ==> floor x = n" by linarith lemma floor_eq3: "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n" @@ -1724,129 +1488,79 @@ lemma floor_eq4: "[| real n \ x; x < real (Suc n) |] ==> nat(floor x) = n" by linarith -lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real(floor r)" - by linarith - -lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real(floor r)" - by linarith - -lemma real_of_int_floor_add_one_ge [simp]: "r \ real(floor r) + 1" +lemma real_of_int_floor_ge_diff_one [simp]: "r - 1 \ real_of_int(floor r)" by linarith -lemma real_of_int_floor_add_one_gt [simp]: "r < real(floor r) + 1" - by linarith - -lemma le_floor: "real a <= x ==> a <= floor x" +lemma real_of_int_floor_gt_diff_one [simp]: "r - 1 < real_of_int(floor r)" by linarith -lemma real_le_floor: "a <= floor x ==> real a <= x" +lemma real_of_int_floor_add_one_ge [simp]: "r \ real_of_int(floor r) + 1" by linarith -lemma le_floor_eq: "(a <= floor x) = (real a <= x)" - by linarith - -lemma floor_less_eq: "(floor x < a) = (x < real a)" +lemma real_of_int_floor_add_one_gt [simp]: "r < real_of_int(floor r) + 1" by linarith -lemma less_floor_eq: "(a < floor x) = (real a + 1 <= x)" - by linarith - -lemma floor_le_eq: "(floor x <= a) = (x < real a + 1)" - by linarith - -lemma floor_eq_iff: "floor x = b \ real b \ x \ x < real (b + 1)" - by linarith +lemma floor_eq_iff: "floor x = b \ of_int b \ x \ x < of_int (b + 1)" +by (simp add: floor_unique_iff) -lemma floor_add [simp]: "floor (x + real a) = floor x + a" - by linarith +lemma floor_add2[simp]: "floor (of_int a + x) = a + floor x" + by (simp add: add.commute) -lemma floor_add2[simp]: "floor (real a + x) = a + floor x" - by linarith - -lemma floor_subtract [simp]: "floor (x - real a) = floor x - a" - by linarith - -lemma floor_divide_real_eq_div: "0 \ b \ floor (a / real b) = floor a div b" +lemma floor_divide_real_eq_div: "0 \ b \ floor (a / real_of_int b) = floor a div b" proof cases assume "0 < b" - { fix i j :: int assume "real i \ a" "a < 1 + real i" - "real j * real b \ a" "a < real b + real j * real b" - then have "i < b + j * b" "j * b < 1 + i" - unfolding real_of_int_less_iff[symmetric] by auto - then have "(j - i div b) * b \ i mod b" "i mod b < ((j - i div b) + 1) * b" + { fix i j :: int assume "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)) + 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" + 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 } + using \0 < b\ unfolding mult_less_cancel_right by auto + } with \0 < b\ show ?thesis by (auto split: floor_split simp: field_simps) qed auto -lemma floor_divide_eq_div: - "floor (real a / real b) = a div b" - using floor_divide_of_int_eq [of a b] real_eq_of_int by simp - lemma floor_divide_eq_div_numeral[simp]: "\numeral a / numeral b::real\ = numeral a div numeral b" - using floor_divide_eq_div[of "numeral a" "numeral b"] by simp + 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" - using floor_divide_eq_div[of "- numeral a" "numeral b"] by simp - -lemma ceiling_real_of_nat [simp]: "ceiling (real (n::nat)) = int n" - by linarith - -lemma real_of_int_ceiling_ge [simp]: "r \ real (ceiling r)" - by linarith - -lemma ceiling_real_of_int [simp]: "ceiling (real (n::int)) = n" - by linarith + by (metis divide_minus_left floor_divide_of_int_eq of_int_neg_numeral of_int_numeral) lemma real_of_int_ceiling_cancel [simp]: - "(real (ceiling x) = x) = (\n::int. x = real n)" - using ceiling_real_of_int by metis + "(real_of_int (ceiling x) = x) = (\n::int. x = real_of_int n)" + using ceiling_of_int by metis -lemma ceiling_eq: "[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1" - by linarith - -lemma ceiling_eq2: "[| real n < x; x \ real n + 1 |] ==> ceiling x = n + 1" - by linarith - -lemma ceiling_eq3: "[| real n - 1 < x; x \ real n |] ==> ceiling x = n" +lemma ceiling_eq: "[| real_of_int n < x; x < real_of_int n + 1 |] ==> ceiling x = n + 1" by linarith -lemma real_of_int_ceiling_diff_one_le [simp]: "real (ceiling r) - 1 \ r" - by linarith - -lemma real_of_int_ceiling_le_add_one [simp]: "real (ceiling r) \ r + 1" +lemma ceiling_eq2: "[| real_of_int n < x; x \ real_of_int n + 1 |] ==> ceiling x = n + 1" by linarith -lemma ceiling_le: "x <= real a ==> ceiling x <= a" +lemma real_of_int_ceiling_diff_one_le [simp]: "real_of_int (ceiling r) - 1 \ r" by linarith -lemma ceiling_le_real: "ceiling x <= a ==> x <= real a" - by linarith - -lemma ceiling_le_eq: "(ceiling x <= a) = (x <= real a)" +lemma real_of_int_ceiling_le_add_one [simp]: "real_of_int (ceiling r) \ r + 1" by linarith -lemma less_ceiling_eq: "(a < ceiling x) = (real a < x)" - by linarith - -lemma ceiling_less_eq: "(ceiling x < a) = (x <= real a - 1)" - by linarith - -lemma le_ceiling_eq: "(a <= ceiling x) = (real a - 1 < x)" +lemma ceiling_le: "x <= real_of_int a ==> ceiling x <= a" by linarith -lemma ceiling_add [simp]: "ceiling (x + real a) = ceiling x + a" +lemma ceiling_le_real: "ceiling x <= a ==> x <= real_of_int a" by linarith -lemma ceiling_subtract [simp]: "ceiling (x - real a) = ceiling x - a" - by linarith - -lemma ceiling_divide_eq_div: "\real a / real b\ = - (- a div b)" - unfolding ceiling_def minus_divide_left real_of_int_minus[symmetric] floor_divide_eq_div by simp_all +lemma ceiling_divide_eq_div: "\real_of_int a / real_of_int b\ = - (- a div b)" + by (metis ceiling_def floor_divide_of_int_eq minus_divide_left of_int_minus) lemma ceiling_divide_eq_div_numeral [simp]: "\numeral a / numeral b :: real\ = - (- numeral a div numeral b)" @@ -1870,13 +1584,12 @@ by (cases "0 <= a & 0 <= b") (auto simp add: nat_mult_distrib[symmetric] nat_mono le_mult_floor) -lemma nat_ceiling_le_eq: "(nat(ceiling x) <= a) = (x <= real a)" +lemma nat_ceiling_le_eq [simp]: "(nat(ceiling x) <= a) = (x <= real a)" by linarith lemma real_nat_ceiling_ge: "x <= real(nat(ceiling x))" by linarith - lemma Rats_no_top_le: "\ q \ \. (x :: real) \ q" by (auto intro!: bexI[of _ "of_nat (nat(ceiling x))"]) linarith @@ -1889,13 +1602,12 @@ subsection \Exponentiation with floor\ lemma floor_power: - assumes "x = real (floor x)" + assumes "x = real_of_int (floor x)" shows "floor (x ^ n) = floor x ^ n" proof - - have *: "x ^ n = real (floor x ^ n)" + have "x ^ n = real_of_int (floor x ^ n)" using assms by (induct n arbitrary: x) simp_all - show ?thesis unfolding real_of_int_inject[symmetric] - unfolding * floor_real_of_int .. + then show ?thesis by linarith qed (* lemma natfloor_power: @@ -1903,12 +1615,13 @@ shows "natfloor (x ^ n) = natfloor x ^ n" proof - from assms have "0 \ floor x" by auto - note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \ floor x`]] + note assms[unfolded natfloor_def of_nat_nat[OF `0 \ floor x`]] from floor_power[OF this] show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \ floor x`, symmetric] by simp qed *) + lemma floor_numeral_power[simp]: "\numeral x ^ n\ = numeral x ^ n" by (metis floor_of_int of_int_numeral of_int_power) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Tue Nov 10 14:18:41 2015 +0000 @@ -341,12 +341,6 @@ lemma of_real_of_int_eq [simp]: "of_real (of_int z) = of_int z" by (cases z rule: int_diff_cases, simp) -lemma of_real_real_of_nat_eq [simp]: "of_real (real n) = of_nat n" - by (simp add: real_of_nat_def) - -lemma of_real_real_of_int_eq [simp]: "of_real (real z) = of_int z" - by (simp add: real_of_int_def) - lemma of_real_numeral [simp]: "of_real (numeral w) = numeral w" using of_real_of_int_eq [of "numeral w"] by simp @@ -997,7 +991,7 @@ also have "\ \ (\i = m * norm (z - w)" - by (simp add: real_of_nat_def) + by simp finally show ?thesis . qed @@ -1291,7 +1285,7 @@ lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" by (simp_all add: dist_norm) - + lemma dist_of_int: "dist (of_int m) (of_int n :: 'a :: real_normed_algebra_1) = of_int \m - n\" proof - have "dist (of_int m) (of_int n :: 'a) = dist (of_int m :: 'a) (of_int m - (of_int (m - n)))" @@ -1300,10 +1294,10 @@ finally show ?thesis . qed -lemma dist_of_nat: +lemma dist_of_nat: "dist (of_nat m) (of_nat n :: 'a :: real_normed_algebra_1) = of_int \int m - int n\" by (subst (1 2) of_int_of_nat_eq [symmetric]) (rule dist_of_int) - + subsection \Bounded Linear and Bilinear Operators\ locale linear = additive f for f :: "'a::real_vector \ 'b::real_vector" + @@ -1571,7 +1565,7 @@ assumes "linear f" "bij f" shows "linear (inv f)" using assms unfolding linear_def linear_axioms_def additive_def by (auto simp: bij_is_surj bij_is_inj surj_f_inv_f intro!: Hilbert_Choice.inv_f_eq) - + instance real_normed_algebra_1 \ perfect_space proof fix x::'a @@ -1737,7 +1731,7 @@ qed next assume "Cauchy f" - show "\e>0. \M. \m\M. \n>m. dist (f m) (f n) < e" + show "\e>0. \M. \m\M. \n>m. dist (f m) (f n) < e" proof (intro allI impI) fix e :: real assume e: "e > 0" with `Cauchy f` obtain M where "\m n. m \ M \ n \ M \ dist (f m) (f n) < e" @@ -1959,8 +1953,7 @@ using ex_le_of_nat[of x] .. note monoD[OF mono this] also have "f (real_of_nat n) \ y" - by (rule LIMSEQ_le_const[OF limseq]) - (auto intro: exI[of _ n] monoD[OF mono] simp: real_eq_of_nat[symmetric]) + by (rule LIMSEQ_le_const[OF limseq]) (auto intro: exI[of _ n] monoD[OF mono]) finally have "f x \ y" . } note le = this have "eventually (\x. real N \ x) at_top" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Series.thy --- a/src/HOL/Series.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Series.thy Tue Nov 10 14:18:41 2015 +0000 @@ -11,7 +11,7 @@ theory Series imports Limits Inequalities -begin +begin subsection \Definition of infinite summability\ @@ -46,7 +46,7 @@ lemma summable_iff_convergent': "summable f \ convergent (\n. setsum f {..n})" - by (simp_all only: summable_iff_convergent convergent_def + by (simp_all only: summable_iff_convergent convergent_def lessThan_Suc_atMost [symmetric] LIMSEQ_Suc_iff[of "\n. setsum f {..n. \in\N. f n = g n" by (auto simp: eventually_at_top_linorder) def C \ "(\kn. setsum f {.. N" @@ -87,7 +87,7 @@ also have "setsum f ... = setsum f {.. {N.. summable f \ (suminf f = x)" by (metis summable_sums sums_summable sums_unique) -lemma summable_sums_iff: +lemma summable_sums_iff: "summable (f :: nat \ 'a :: {comm_monoid_add,t2_space}) \ f sums suminf f" by (auto simp: sums_iff summable_sums) @@ -351,7 +351,7 @@ lemma suminf_split_head: "summable f \ (\n. f (Suc n)) = suminf f - f 0" using suminf_split_initial_segment[of 1] by simp -lemma suminf_exist_split: +lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable f" shows "\N. \n\N. norm (\i. f (i + n)) < r" proof - @@ -428,11 +428,10 @@ { assume "c \ 0" hence "filterlim (\n. of_nat n * norm c) at_top sequentially" - unfolding real_of_nat_def[symmetric] by (subst mult.commute) (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_real_sequentially) hence "\convergent (\n. norm (\ksummable (\_. c)" unfolding summable_iff_convergent using convergent_norm by blast } @@ -529,7 +528,7 @@ by (auto simp: eventually_at_top_linorder) thus "norm c < 1" using one_le_power[of "norm c" n] by (cases "norm c \ 1") (linarith, simp) qed (rule summable_geometric) - + end lemma power_half_series: "(\n. (1/2::real)^Suc n) sums 1" @@ -638,7 +637,7 @@ subsection \The Ratio Test\ -lemma summable_ratio_test: +lemma summable_ratio_test: assumes "c < 1" "\n. n \ N \ norm (f (Suc n)) \ c * norm (f n)" shows "summable f" proof cases @@ -681,7 +680,7 @@ shows "summable (\n. norm (a n) * r^n)" proof (rule summable_comparison_test') show "summable (\n. M * (r / r0) ^ n)" - using assms + using assms by (auto simp add: summable_mult summable_geometric) next fix n @@ -815,7 +814,7 @@ lemma summable_zero_power' [simp]: "summable (\n. f n * 0 ^ n :: 'a :: {ring_1,topological_space})" proof - - have "(\n. f n * 0 ^ n :: 'a) = (\n. if n = 0 then f 0 * 0^0 else 0)" + have "(\n. f n * 0 ^ n :: 'a) = (\n. if n = 0 then f 0 * 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp @@ -845,7 +844,7 @@ have "summable (\n. f (Suc n) * z ^ n) \ summable (\n. f (Suc n) * z ^ Suc n)" proof assume "summable (\n. f (Suc n) * z ^ n)" - from summable_mult2[OF this, of z] show "summable (\n. f (Suc n) * z ^ Suc n)" + from summable_mult2[OF this, of z] show "summable (\n. f (Suc n) * z ^ Suc n)" by (simp add: power_commutes algebra_simps) next assume "summable (\n. f (Suc n) * z ^ Suc n)" @@ -864,7 +863,7 @@ proof - from assms show "summable (\n. f (Suc n) * z ^ n)" by (subst summable_powser_split_head) - from suminf_mult2[OF this, of z] + from suminf_mult2[OF this, of z] have "(\n. f (Suc n) * z ^ n) * z = (\n. f (Suc n) * z ^ Suc n)" by (simp add: power_commutes algebra_simps) also from assms have "\ = suminf (\n. f n * z ^ n) - f 0" @@ -878,9 +877,9 @@ assumes summable: "summable f" and e: "e > (0::real)" obtains N where "\m n. m \ N \ norm (\k=m..n. f k) < e" proof - - from summable have "Cauchy (\n. \kn. \km n. m \ N \ n \ N \ norm ((\kk N" @@ -896,10 +895,10 @@ thus ?thesis by (rule that) qed -lemma powser_sums_if: +lemma powser_sums_if: "(\n. (if n = m then (1 :: 'a :: {ring_1,topological_space}) else 0) * z^n) sums z^m" proof - - have "(\n. (if n = m then 1 else 0) * z^n) = (\n. if n = m then z^n else 0)" + have "(\n. (if n = m then 1 else 0) * z^n) = (\n. if n = m then z^n else 0)" by (intro ext) auto thus ?thesis by (simp add: sums_single) qed @@ -918,7 +917,7 @@ have smaller: "\n. (\i g) i) \ suminf f" proof fix n - have "\ n' \ (g ` {.. n' \ (g ` {..n'. n' < n \ g n' < m" by blast @@ -927,7 +926,7 @@ also have "\ \ (\i \ suminf f" - using \summable f\ + using \summable f\ by (rule setsum_le_suminf) (simp add: pos) finally show "(\i g) i) \ suminf f" by simp qed @@ -991,9 +990,9 @@ from filterlim_subseq[OF subseq] have g_inv_ex: "\m. g m \ n" for n by (auto simp: filterlim_at_top eventually_at_top_linorder) hence g_inv: "g (g_inv n) \ n" for n unfolding g_inv_def by (rule LeastI_ex) - have g_inv_least: "m \ g_inv n" if "g m \ n" for m n using that + have g_inv_least: "m \ g_inv n" if "g m \ n" for m n using that unfolding g_inv_def by (rule Least_le) - have g_inv_least': "g m < n" if "m < g_inv n" for m n using that g_inv_least[of n m] by linarith + have g_inv_least': "g m < n" if "m < g_inv n" for m n using that g_inv_least[of n m] by linarith have "(\n. \kn. \kkk\g`{.. = (\k = (\kkk g (g_inv n)" by (rule g_inv) finally have "K \ g_inv n" using subseq by (simp add: strict_mono_less_eq subseq_strict_mono) } - hence "filterlim g_inv at_top sequentially" + hence "filterlim g_inv at_top sequentially" by (auto simp: filterlim_at_top eventually_at_top_linorder) from lim and this have "(\n. \k c" by (rule filterlim_compose) finally show "(\n. \k c" . @@ -1030,7 +1029,7 @@ shows "summable (\n. f (g n)) \ summable f" using sums_mono_reindex[of g f, OF assms] by (simp add: summable_def) -lemma suminf_mono_reindex: +lemma suminf_mono_reindex: assumes "subseq g" "\n. n \ range g \ f n = (0 :: 'a :: {t2_space,comm_monoid_add})" shows "suminf (\n. f (g n)) = suminf f" proof (cases "summable f") @@ -1042,7 +1041,7 @@ hence "\(\c. (\n. f (g n)) sums c)" unfolding summable_def by blast hence "suminf (\n. f (g n)) = The (\_. False)" by (simp add: suminf_def) ultimately show ?thesis by simp -qed (insert sums_mono_reindex[of g f, OF assms] summable_mono_reindex[of g f, OF assms], +qed (insert sums_mono_reindex[of g f, OF assms] summable_mono_reindex[of g f, OF assms], simp_all add: sums_iff) end diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/TPTP/THF_Arith.thy --- a/src/HOL/TPTP/THF_Arith.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/TPTP/THF_Arith.thy Tue Nov 10 14:18:41 2015 +0000 @@ -56,7 +56,7 @@ overloading int_to_real \ "to_real :: int \ real" begin - definition "int_to_real (n::int) = real n" + definition "int_to_real (n::int) = real_of_int n" end overloading rat_to_real \ "to_real :: rat \ real" @@ -79,7 +79,7 @@ by (metis int_to_rat_def rat_is_int_def) lemma to_real_is_int [intro, simp]: "is_int (to_real (n::int))" -by (metis Ints_real_of_int int_to_real_def real_is_int_def) +by (metis Ints_of_int int_to_real_def real_is_int_def) lemma to_real_is_rat [intro, simp]: "is_rat (to_real (q::rat))" by (metis Rats_of_rat rat_to_real_def real_is_rat_def) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Tools/SMT/z3_real.ML --- a/src/HOL/Tools/SMT/z3_real.ML Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Tools/SMT/z3_real.ML Tue Nov 10 14:18:41 2015 +0000 @@ -13,14 +13,14 @@ fun real_term_parser (SMTLIB.Dec (i, 0), []) = SOME (HOLogic.mk_number @{typ Real.real} i) | real_term_parser (SMTLIB.Sym "/", [t1, t2]) = SOME (@{term "Rings.divide :: real => _"} $ t1 $ t2) - | real_term_parser (SMTLIB.Sym "to_real", [t]) = SOME (@{term "Real.real :: int => _"} $ t) + | real_term_parser (SMTLIB.Sym "to_real", [t]) = SOME (@{term "Int.of_int :: int => _"} $ t) | real_term_parser _ = NONE fun abstract abs t = (case t of (c as @{term "Rings.divide :: real => _"}) $ t1 $ t2 => abs t1 ##>> abs t2 #>> (fn (u1, u2) => SOME (c $ u1 $ u2)) - | (c as @{term "Real.real :: int => _"}) $ t => + | (c as @{term "Int.of_int :: int => _"}) $ t => abs t #>> (fn u => SOME (c $ u)) | _ => pair NONE) diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/Transcendental.thy Tue Nov 10 14:18:41 2015 +0000 @@ -11,10 +11,11 @@ begin lemma reals_Archimedean4: - assumes "0 < y" "0 \ x" + assumes "0 < y" "0 \ x" obtains n where "real n * y \ x" "x < real (Suc n) * y" using floor_correct [of "x/y"] assms - by (auto simp: Real.real_of_nat_Suc field_simps intro: that [of "nat (floor (x/y))"]) + apply (auto simp: field_simps intro!: that [of "nat (floor (x/y))"]) +by (metis (no_types, hide_lams) divide_inverse divide_self_if floor_correct mult.left_commute mult.right_neutral mult_eq_0_iff order_refl order_trans real_mult_le_cancel_iff2) lemma fact_in_Reals: "fact n \ \" by (induction n) auto @@ -24,11 +25,8 @@ lemma of_real_fact [simp]: "of_real (fact n) = fact n" by (metis of_nat_fact of_real_of_nat_eq) -lemma real_fact_nat [simp]: "real (fact n :: nat) = fact n" - by (simp add: real_of_nat_def) - -lemma real_fact_int [simp]: "real (fact n :: int) = fact n" - by (metis of_int_of_nat_eq of_nat_fact real_of_int_def) +lemma of_int_fact [simp]: "of_int (fact n :: int) = fact n" + by (metis of_int_of_nat_eq of_nat_fact) lemma norm_fact [simp]: "norm (fact n :: 'a :: {real_normed_algebra_1}) = fact n" @@ -155,26 +153,26 @@ moreover obtain N where N: "norm x / (1 - norm x) < of_int N" using ex_le_of_int by (meson ex_less_of_int) - ultimately have N0: "N>0" + ultimately have N0: "N>0" by auto - then have *: "real (N + 1) * norm x / real N < 1" + then have *: "real_of_int (N + 1) * norm x / real_of_int N < 1" using N assms by (auto simp: field_simps) { fix n::nat assume "N \ int n" - then have "real N * real_of_nat (Suc n) \ real_of_nat n * real (1 + N)" + then have "real_of_int N * real_of_nat (Suc n) \ real_of_nat n * real_of_int (1 + N)" by (simp add: algebra_simps) - then have "(real N * real_of_nat (Suc n)) * (norm x * norm (x ^ n)) - \ (real_of_nat n * real (1 + N)) * (norm x * norm (x ^ n))" + then have "(real_of_int N * real_of_nat (Suc n)) * (norm x * norm (x ^ n)) + \ (real_of_nat n * (1 + N)) * (norm x * norm (x ^ n))" using N0 mult_mono by fastforce - then have "real N * (norm x * (real_of_nat (Suc n) * norm (x ^ n))) - \ real_of_nat n * (norm x * (real (1 + N) * norm (x ^ n)))" + then have "real_of_int N * (norm x * (real_of_nat (Suc n) * norm (x ^ n))) + \ real_of_nat n * (norm x * ((1 + N) * norm (x ^ n)))" by (simp add: algebra_simps) } note ** = this show ?thesis using * apply (rule summable_LIMSEQ_zero [OF summable_ratio_test, where N1="nat N"]) - apply (simp add: N0 norm_mult field_simps ** - del: of_nat_Suc real_of_int_add) + apply (simp add: N0 norm_mult field_simps ** + del: of_nat_Suc of_int_add) done qed @@ -184,16 +182,6 @@ using powser_times_n_limit_0 [of "inverse x"] by (simp add: norm_divide divide_simps) -lemma lim_1_over_n: "((\n. 1 / of_nat n) ---> (0::'a::real_normed_field)) sequentially" - apply (clarsimp simp: lim_sequentially norm_divide dist_norm divide_simps) - apply (auto simp: mult_ac dest!: ex_less_of_nat_mult [of _ 1]) - by (metis le_eq_less_or_eq less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono - of_nat_less_0_iff of_nat_less_iff zero_less_mult_iff zero_less_one) - -lemma lim_inverse_n: "((\n. inverse(of_nat n)) ---> (0::'a::real_normed_field)) sequentially" - using lim_1_over_n - by (simp add: inverse_eq_divide) - lemma sum_split_even_odd: fixes f :: "nat \ real" shows @@ -751,7 +739,7 @@ then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ n)" using False summable_mult2 [of "\n. (of_nat (Suc n) * c(Suc n) * x ^ n) * x" "inverse x"] by (simp add: mult.assoc) (auto simp: ac_simps) - then show ?thesis + then show ?thesis by (simp add: diffs_def) qed @@ -798,14 +786,14 @@ shows "((\x. \n. c n * x^n) has_field_derivative (\n. diffs c n * x^n)) (at x)" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force simp del: of_real_add) - + lemma isCont_powser: fixes K x :: "'a::{real_normed_field,banach}" assumes "summable (\n. c n * K ^ n)" assumes "norm x < norm K" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms] by (blast intro!: DERIV_isCont) - + lemmas isCont_powser' = isCont_o2[OF _ isCont_powser] lemma isCont_powser_converges_everywhere: @@ -815,7 +803,7 @@ using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force intro!: DERIV_isCont simp del: of_real_add) -lemma powser_limit_0: +lemma powser_limit_0: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. norm x < s \ (\n. a n * x ^ n) sums (f x)" @@ -835,16 +823,16 @@ by (blast intro: DERIV_continuous) then have "((\x. \n. a n * x ^ n) ---> a 0) (at 0)" by (simp add: continuous_within powser_zero) - then show ?thesis + then show ?thesis apply (rule Lim_transform) apply (auto simp add: LIM_eq) apply (rule_tac x="s" in exI) - using s + using s apply (auto simp: sm [THEN sums_unique]) done qed -lemma powser_limit_0_strong: +lemma powser_limit_0_strong: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. x \ 0 \ norm x < s \ (\n. a n * x ^ n) sums (f x)" @@ -906,7 +894,7 @@ from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \0 < ?r\, unfolded real_norm_def] obtain s where s_bound: "0 < s \ (\x. x \ 0 \ \x\ < s \ \?diff n x - f' x0 n\ < ?r)" by auto - have "0 < ?s n" by (rule someI2[where a=s]) (auto simp add: s_bound simp del: real_of_nat_Suc) + have "0 < ?s n" by (rule someI2[where a=s]) (auto simp add: s_bound simp del: of_nat_Suc) thus "0 < x" unfolding \x = ?s n\ . qed qed auto @@ -964,9 +952,8 @@ qed auto also have "\ = of_nat (card {.. = real ?N * ?r" - unfolding real_eq_of_nat by auto - also have "\ = r/3" by (auto simp del: real_of_nat_Suc) + also have "\ = real ?N * ?r" by simp + also have "\ = r/3" by (auto simp del: of_nat_Suc) finally have "\\n < r / 3" (is "?diff_part < r / 3") . from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] @@ -1050,9 +1037,10 @@ unfolding power_add[symmetric] using \p \ n\ by auto qed also have "\ = real (Suc n) * R' ^ n" - unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto + unfolding setsum_constant card_atLeastLessThan by auto finally show "\\p \ \real (Suc n)\ * \R' ^ n\" - unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \0 < R'\]]] . + unfolding abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \0 < R'\]]] + by linarith show "0 \ \f n\ * \x - y\" unfolding abs_mult[symmetric] by auto qed @@ -1064,7 +1052,7 @@ { fix n show "DERIV (\ x. ?f x n) x0 :> (?f' x0 n)" - by (auto intro!: derivative_eq_intros simp del: power_Suc simp: real_of_nat_def) + by (auto intro!: derivative_eq_intros simp del: power_Suc) } { fix x @@ -1112,7 +1100,7 @@ lemma isCont_pochhammer [continuous_intros]: "isCont (\z::'a::real_normed_field. pochhammer z n) z" by (induction n) (auto intro!: continuous_intros simp: pochhammer_rec') -lemma continuous_on_pochhammer [continuous_intros]: +lemma continuous_on_pochhammer [continuous_intros]: fixes A :: "'a :: real_normed_field set" shows "continuous_on A (\z. pochhammer z n)" by (intro continuous_at_imp_continuous_on ballI isCont_pochhammer) @@ -1133,7 +1121,7 @@ obtain r :: real where r0: "0 < r" and r1: "r < 1" using dense [OF zero_less_one] by fast obtain N :: nat where N: "norm x < real N * r" - using ex_less_of_nat_mult r0 real_of_nat_def by auto + using ex_less_of_nat_mult r0 by auto from r1 show ?thesis proof (rule summable_ratio_test [rule_format]) fix n :: nat @@ -1164,7 +1152,7 @@ by (simp add: norm_power_ineq) qed -lemma summable_exp: +lemma summable_exp: fixes x :: "'a::{real_normed_field,banach}" shows "summable (\n. inverse (fact n) * x^n)" using summable_exp_generic [where x=x] @@ -1190,7 +1178,7 @@ apply (simp del: of_real_add) done -declare DERIV_exp[THEN DERIV_chain2, derivative_intros] +declare DERIV_exp[THEN DERIV_chain2, derivative_intros] DERIV_exp[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemma norm_exp: "norm (exp x) \ exp (norm x)" @@ -1279,7 +1267,7 @@ (\i\Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) = (\i\Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n-i)))" by (simp only: setsum.distrib [symmetric] scaleR_left_distrib [symmetric] - real_of_nat_add [symmetric]) simp + of_nat_add [symmetric]) simp also have "\ = real (Suc n) *\<^sub>R (\i\Suc n. S x i * S y (Suc n-i))" by (simp only: scaleR_right.setsum) finally show @@ -1338,7 +1326,7 @@ by (induct n) (auto simp add: distrib_left exp_add mult.commute) corollary exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" - by (simp add: exp_of_nat_mult real_of_nat_def) + by (simp add: exp_of_nat_mult) lemma exp_setsum: "finite I \ exp(setsum f I) = setprod (\x. exp(f x)) I" by (induction I rule: finite_induct) (auto simp: exp_add_commuting mult.commute) @@ -1480,7 +1468,7 @@ definition ln_real :: "real \ real" where "ln_real x = (THE u. exp u = x)" -instance +instance by intro_classes (simp add: ln_real_def) end @@ -1488,106 +1476,106 @@ lemma powr_eq_0_iff [simp]: "w powr z = 0 \ w = 0" by (simp add: powr_def) -lemma ln_exp [simp]: +lemma ln_exp [simp]: fixes x::real shows "ln (exp x) = x" by (simp add: ln_real_def) -lemma exp_ln [simp]: +lemma exp_ln [simp]: fixes x::real shows "0 < x \ exp (ln x) = x" by (auto dest: exp_total) -lemma exp_ln_iff [simp]: +lemma exp_ln_iff [simp]: fixes x::real shows "exp (ln x) = x \ 0 < x" by (metis exp_gt_zero exp_ln) -lemma ln_unique: +lemma ln_unique: fixes x::real shows "exp y = x \ ln x = y" by (erule subst, rule ln_exp) -lemma ln_mult: +lemma ln_mult: fixes x::real shows "0 < x \ 0 < y \ ln (x * y) = ln x + ln y" by (rule ln_unique) (simp add: exp_add) -lemma ln_setprod: - fixes f:: "'a => real" +lemma ln_setprod: + fixes f:: "'a => real" shows "\finite I; \i. i \ I \ f i > 0\ \ ln(setprod f I) = setsum (\x. ln(f x)) I" by (induction I rule: finite_induct) (auto simp: ln_mult setprod_pos) -lemma ln_inverse: +lemma ln_inverse: fixes x::real shows "0 < x \ ln (inverse x) = - ln x" by (rule ln_unique) (simp add: exp_minus) -lemma ln_div: +lemma ln_div: fixes x::real shows "0 < x \ 0 < y \ ln (x / y) = ln x - ln y" by (rule ln_unique) (simp add: exp_diff) lemma ln_realpow: "0 < x \ ln (x^n) = real n * ln x" by (rule ln_unique) (simp add: exp_real_of_nat_mult) -lemma ln_less_cancel_iff [simp]: +lemma ln_less_cancel_iff [simp]: fixes x::real shows "0 < x \ 0 < y \ ln x < ln y \ x < y" by (subst exp_less_cancel_iff [symmetric]) simp -lemma ln_le_cancel_iff [simp]: +lemma ln_le_cancel_iff [simp]: fixes x::real shows "0 < x \ 0 < y \ ln x \ ln y \ x \ y" by (simp add: linorder_not_less [symmetric]) -lemma ln_inj_iff [simp]: +lemma ln_inj_iff [simp]: fixes x::real shows "0 < x \ 0 < y \ ln x = ln y \ x = y" by (simp add: order_eq_iff) -lemma ln_add_one_self_le_self [simp]: +lemma ln_add_one_self_le_self [simp]: fixes x::real shows "0 \ x \ ln (1 + x) \ x" apply (rule exp_le_cancel_iff [THEN iffD1]) apply (simp add: exp_ge_add_one_self_aux) done -lemma ln_less_self [simp]: +lemma ln_less_self [simp]: fixes x::real shows "0 < x \ ln x < x" by (rule order_less_le_trans [where y="ln (1 + x)"]) simp_all -lemma ln_ge_zero [simp]: +lemma ln_ge_zero [simp]: fixes x::real shows "1 \ x \ 0 \ ln x" using ln_le_cancel_iff [of 1 x] by simp -lemma ln_ge_zero_imp_ge_one: +lemma ln_ge_zero_imp_ge_one: fixes x::real shows "0 \ ln x \ 0 < x \ 1 \ x" using ln_le_cancel_iff [of 1 x] by simp -lemma ln_ge_zero_iff [simp]: +lemma ln_ge_zero_iff [simp]: fixes x::real shows "0 < x \ 0 \ ln x \ 1 \ x" using ln_le_cancel_iff [of 1 x] by simp -lemma ln_less_zero_iff [simp]: +lemma ln_less_zero_iff [simp]: fixes x::real shows "0 < x \ ln x < 0 \ x < 1" using ln_less_cancel_iff [of x 1] by simp -lemma ln_gt_zero: +lemma ln_gt_zero: fixes x::real shows "1 < x \ 0 < ln x" using ln_less_cancel_iff [of 1 x] by simp -lemma ln_gt_zero_imp_gt_one: +lemma ln_gt_zero_imp_gt_one: fixes x::real shows "0 < ln x \ 0 < x \ 1 < x" using ln_less_cancel_iff [of 1 x] by simp -lemma ln_gt_zero_iff [simp]: +lemma ln_gt_zero_iff [simp]: fixes x::real shows "0 < x \ 0 < ln x \ 1 < x" using ln_less_cancel_iff [of 1 x] by simp -lemma ln_eq_zero_iff [simp]: +lemma ln_eq_zero_iff [simp]: fixes x::real shows "0 < x \ ln x = 0 \ x = 1" using ln_inj_iff [of x 1] by simp -lemma ln_less_zero: +lemma ln_less_zero: fixes x::real shows "0 < x \ x < 1 \ ln x < 0" by simp -lemma ln_neg_is_const: +lemma ln_neg_is_const: fixes x::real shows "x \ 0 \ ln x = (THE x. False)" by (auto simp add: ln_real_def intro!: arg_cong[where f=The]) -lemma isCont_ln: +lemma isCont_ln: fixes x::real assumes "x \ 0" shows "isCont ln x" proof cases assume "0 < x" @@ -1603,7 +1591,7 @@ intro!: exI[of _ "\x\"]) qed -lemma tendsto_ln [tendsto_intros]: +lemma tendsto_ln [tendsto_intros]: fixes a::real shows "(f ---> a) F \ a \ 0 \ ((\x. ln (f x)) ---> ln a) F" by (rule isCont_tendsto_compose [OF isCont_ln]) @@ -1709,9 +1697,9 @@ have "(2::nat) * 2 ^ n \ fact (n + 2)" by (induct n) simp_all hence "real ((2::nat) * 2 ^ n) \ real_of_nat (fact (n + 2))" - by (simp only: real_of_nat_le_iff) + by (simp only: of_nat_le_iff) hence "((2::real) * 2 ^ n) \ fact (n + 2)" - unfolding of_nat_fact real_of_nat_def + unfolding of_nat_fact by (simp add: of_nat_mult of_nat_power) hence "inverse (fact (n + 2)) \ inverse ((2::real) * 2 ^ n)" by (rule le_imp_inverse_le) simp @@ -1843,7 +1831,7 @@ qed lemma ln_one_minus_pos_lower_bound: - fixes x::real + fixes x::real shows "0 <= x \ x <= (1 / 2) \ - x - 2 * x\<^sup>2 <= ln (1 - x)" proof - assume a: "0 <= x" and b: "x <= (1 / 2)" @@ -1973,13 +1961,13 @@ fixes x::real shows "0 < x \ ln x \ x - 1" using exp_ge_add_one_self[of "ln x"] by simp -corollary ln_diff_le: - fixes x::real +corollary ln_diff_le: + fixes x::real shows "0 < x \ 0 < y \ ln x - ln y \ (x - y) / y" by (simp add: ln_div [symmetric] diff_divide_distrib ln_le_minus_one) lemma ln_eq_minus_one: - fixes x::real + fixes x::real assumes "0 < x" "ln x = x - 1" shows "x = 1" proof - @@ -2328,7 +2316,7 @@ lemma le_log_iff: assumes "1 < b" "x > 0" shows "y \ log b x \ b powr y \ (x::real)" - using assms + using assms apply auto apply (metis (no_types, hide_lams) less_irrefl less_le_trans linear powr_le_cancel_iff powr_log_cancel zero_less_one) @@ -2358,10 +2346,10 @@ by (auto simp add: floor_eq_iff powr_le_iff less_powr_iff) lemma powr_realpow: "0 < x ==> x powr (real n) = x^n" - by (induct n) (simp_all add: ac_simps powr_add real_of_nat_Suc) + by (induct n) (simp_all add: ac_simps powr_add of_nat_Suc) lemma powr_realpow_numeral: "0 < x \ x powr (numeral n :: real) = x ^ (numeral n)" - unfolding real_of_nat_numeral [symmetric] by (rule powr_realpow) + by (metis of_nat_numeral powr_realpow) lemma powr2_sqrt[simp]: "0 < x \ sqrt x powr 2 = x" by(simp add: powr_realpow_numeral) @@ -2517,7 +2505,7 @@ qed lemma tendsto_powr [tendsto_intros]: - fixes a::real + fixes a::real assumes f: "(f ---> a) F" and g: "(g ---> b) F" and a: "a \ 0" shows "((\x. f x powr g x) ---> a powr b) F" unfolding powr_def @@ -2561,7 +2549,7 @@ next { assume "a = 0" with f f_nonneg have "LIM x inf F (principal {x. f x \ 0}). f x :> at_right 0" - by (auto simp add: filterlim_at eventually_inf_principal le_less + by (auto simp add: filterlim_at eventually_inf_principal le_less elim: eventually_elim1 intro: tendsto_mono inf_le1) then have "((\x. exp (g x * ln (f x))) ---> 0) (inf F (principal {x. f x \ 0}))" by (auto intro!: filterlim_compose[OF exp_at_bot] filterlim_compose[OF ln_at_0] @@ -2770,10 +2758,10 @@ by (metis Reals_cases Reals_of_real cos_of_real) lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff" - by (simp add: diffs_def sin_coeff_Suc real_of_nat_def del: of_nat_Suc) + by (simp add: diffs_def sin_coeff_Suc del: of_nat_Suc) lemma diffs_cos_coeff: "diffs cos_coeff = (\n. - sin_coeff n)" - by (simp add: diffs_def cos_coeff_Suc real_of_nat_def del: of_nat_Suc) + by (simp add: diffs_def cos_coeff_Suc del: of_nat_Suc) text\Now at last we can get the derivatives of exp, sin and cos\ @@ -2884,7 +2872,7 @@ lemma DERIV_fun_cos: "DERIV g x :> m \ DERIV (\x. cos(g x)) x :> -sin(g x) * m" - by (auto intro!: derivative_eq_intros simp: real_of_nat_def) + by (auto intro!: derivative_eq_intros) subsection \Deriving the Addition Formulas\ @@ -2903,7 +2891,7 @@ have "(cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) = (if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" using \n\p\ - by (auto simp: * algebra_simps cos_coeff_def binomial_fact real_of_nat_def) + by (auto simp: * algebra_simps cos_coeff_def binomial_fact) } then have "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = @@ -2942,7 +2930,7 @@ (if even p \ odd n then -((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" using \n\p\ - by (auto simp: algebra_simps sin_coeff_def binomial_fact real_of_nat_def) + by (auto simp: algebra_simps sin_coeff_def binomial_fact) } then have "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = @@ -2977,7 +2965,7 @@ else 0)" by (auto simp: setsum_right_distrib field_simps scaleR_conv_of_real nonzero_of_real_divide) also have "... = cos_coeff p *\<^sub>R ((x + y) ^ p)" - by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real real_of_nat_def atLeast0AtMost) + by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real atLeast0AtMost) finally have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = cos_coeff p *\<^sub>R ((x + y) ^ p)" . @@ -3114,7 +3102,7 @@ lemma DERIV_fun_pow: "DERIV g x :> m ==> DERIV (\x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" - by (auto intro!: derivative_eq_intros simp: real_of_nat_def) + by (auto intro!: derivative_eq_intros simp:) lemma DERIV_fun_exp: "DERIV g x :> m ==> DERIV (\x. exp(g x)) x :> exp(g x) * m" @@ -3155,7 +3143,7 @@ hence "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)" by (intro mult_strict_right_mono zero_less_power \0 < x\) thus "0 < ?f n" - by (simp add: real_of_nat_def divide_simps mult_ac del: mult_Suc) + by (simp add: divide_simps mult_ac del: mult_Suc) qed have sums: "?f sums sin x" by (rule sin_paired [THEN sums_group], simp) @@ -3183,7 +3171,7 @@ lemmas realpow_num_eq_if = power_eq_if -lemma sumr_pos_lt_pair: +lemma sumr_pos_lt_pair: fixes f :: "nat \ real" shows "\summable f; \d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\ @@ -3212,13 +3200,12 @@ { fix d let ?six4d = "Suc (Suc (Suc (Suc (Suc (Suc (4 * d))))))" have "(4::real) * (fact (?six4d)) < (Suc (Suc (?six4d)) * fact (Suc (?six4d)))" - unfolding real_of_nat_mult - by (rule mult_strict_mono) (simp_all add: fact_less_mono) + unfolding of_nat_mult by (rule mult_strict_mono) (simp_all add: fact_less_mono) then have "(4::real) * (fact (?six4d)) < (fact (Suc (Suc (?six4d))))" - by (simp only: fact.simps(2) [of "Suc (?six4d)"] real_of_nat_def of_nat_mult of_nat_fact) + by (simp only: fact.simps(2) [of "Suc (?six4d)"] of_nat_mult of_nat_fact) then have "(4::real) * inverse (fact (Suc (Suc (?six4d)))) < inverse (fact (?six4d))" by (simp add: inverse_eq_divide less_divide_eq) - } + } then show ?thesis by (force intro!: sumr_pos_lt_pair [OF sm] simp add: divide_inverse algebra_simps) qed @@ -3367,6 +3354,9 @@ lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" by (simp add: cos_add) +lemma cos_periodic_pi2 [simp]: "cos (pi + x) = - cos x" + by (simp add: cos_add) + lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x" by (simp add: sin_add sin_double cos_double) @@ -3374,13 +3364,13 @@ by (simp add: cos_add sin_double cos_double) lemma cos_npi [simp]: "cos (real n * pi) = (- 1) ^ n" - by (induct n) (auto simp: real_of_nat_Suc distrib_right) + by (induct n) (auto simp: distrib_right) lemma cos_npi2 [simp]: "cos (pi * real n) = (- 1) ^ n" by (metis cos_npi mult.commute) lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0" - by (induct n) (auto simp: real_of_nat_Suc distrib_right) + by (induct n) (auto simp: of_nat_Suc distrib_right) lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0" by (simp add: mult.commute [of pi]) @@ -3419,33 +3409,33 @@ apply (simp add: field_simps) done -lemma sin_diff_sin: +lemma sin_diff_sin: fixes w :: "'a::{real_normed_field,banach,field}" shows "sin(w) - sin(z) = 2 * sin((w - z) / 2) * cos((w + z) / 2)" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done -lemma cos_plus_cos: +lemma cos_plus_cos: fixes w :: "'a::{real_normed_field,banach,field}" shows "cos(w) + cos(z) = 2 * cos((w + z) / 2) * cos((w - z) / 2)" apply (simp add: mult.assoc cos_times_cos) apply (simp add: field_simps) done -lemma cos_diff_cos: +lemma cos_diff_cos: fixes w :: "'a::{real_normed_field,banach,field}" shows "cos(w) - cos(z) = 2 * sin((w + z) / 2) * sin((z - w) / 2)" apply (simp add: mult.assoc sin_times_sin) apply (simp add: field_simps) done -lemma cos_double_cos: +lemma cos_double_cos: fixes z :: "'a::{real_normed_field,banach}" shows "cos(2 * z) = 2 * cos z ^ 2 - 1" by (simp add: cos_double sin_squared_eq) -lemma cos_double_sin: +lemma cos_double_sin: fixes z :: "'a::{real_normed_field,banach}" shows "cos(2 * z) = 1 - 2 * sin z ^ 2" by (simp add: cos_double sin_squared_eq) @@ -3466,7 +3456,7 @@ by (metis sin_periodic_pi2 add_diff_eq mult_2 sin_pi_minus) lemma cos_2pi_minus [simp]: "cos (2*pi - x) = cos x" - by (metis (no_types, hide_lams) cos_add cos_minus cos_two_pi sin_minus sin_two_pi + by (metis (no_types, hide_lams) cos_add cos_minus cos_two_pi sin_minus sin_two_pi diff_0_right minus_diff_eq mult_1 mult_zero_left uminus_add_conv_diff) lemma sin_gt_zero2: "\0 < x; x < pi/2\ \ 0 < sin x" @@ -3566,13 +3556,6 @@ done qed -lemma reals_Archimedean4': - "\0 < y; 0 \ x\ \ \n. real n * y \ x \ x < real (Suc n) * y" -apply (rule_tac x="nat (floor (x/y))" in exI) -using floor_correct [of "x/y"] -apply (auto simp: Real.real_of_nat_Suc field_simps) -done - lemma cos_zero_lemma: "\0 \ x; cos x = 0\ \ \n::nat. odd n & x = real n * (pi/2)" @@ -3580,7 +3563,7 @@ apply (auto simp: ) apply (subgoal_tac "0 \ x - real n * pi & (x - real n * pi) \ pi & (cos (x - real n * pi) = 0) ") -apply (auto simp: algebra_simps real_of_nat_Suc) +apply (auto simp: algebra_simps of_nat_Suc) prefer 2 apply (simp add: cos_diff) apply (simp add: cos_diff) apply (subgoal_tac "EX! x. 0 \ x & x \ pi & cos x = 0") @@ -3589,7 +3572,7 @@ apply (drule_tac x = "pi/2" in spec) apply (simp add: cos_diff) apply (rule_tac x = "Suc (2 * n)" in exI) -apply (simp add: real_of_nat_Suc algebra_simps, auto) +apply (simp add: of_nat_Suc algebra_simps, auto) done lemma sin_zero_lemma: @@ -3597,7 +3580,7 @@ \n::nat. even n & x = real n * (pi/2)" apply (subgoal_tac "\n::nat. ~ even n & x + pi/2 = real n * (pi/2) ") apply (clarify, rule_tac x = "n - 1" in exI) - apply (auto elim!: oddE simp add: real_of_nat_Suc field_simps)[1] + apply (auto elim!: oddE simp add: of_nat_Suc field_simps)[1] apply (rule cos_zero_lemma) apply (auto simp: cos_add) done @@ -3611,7 +3594,7 @@ assume "odd n" then obtain m where "n = 2 * m + 1" .. then have "cos (real n * pi / 2) = 0" - by (simp add: field_simps real_of_nat_Suc) (simp add: cos_add add_divide_distrib) + by (simp add: field_simps of_nat_Suc) (simp add: cos_add add_divide_distrib) } note * = this show ?thesis apply (rule iffI) @@ -3637,18 +3620,18 @@ lemma cos_zero_iff_int: - "cos x = 0 \ (\n::int. odd n & x = real n * (pi/2))" + "cos x = 0 \ (\n::int. odd n \ x = of_int n * (pi/2))" proof safe assume "cos x = 0" - then show "\n::int. odd n & x = real n * (pi/2)" + then show "\n::int. odd n & x = of_int n * (pi/2)" apply (simp add: cos_zero_iff, safe) - apply (metis even_int_iff real_of_int_of_nat_eq) + apply (metis even_int_iff of_int_of_nat_eq) apply (rule_tac x="- (int n)" in exI, simp) done next fix n::int assume "odd n" - then show "cos (real n * (pi / 2)) = 0" + then show "cos (of_int n * (pi / 2)) = 0" apply (simp add: cos_zero_iff) apply (case_tac n rule: int_cases2, simp) apply (rule disjI2) @@ -3657,18 +3640,18 @@ qed lemma sin_zero_iff_int: - "sin x = 0 \ (\n::int. even n & (x = real n * (pi/2)))" + "sin x = 0 \ (\n::int. even n & (x = of_int n * (pi/2)))" proof safe assume "sin x = 0" - then show "\n::int. even n \ x = real n * (pi / 2)" + then show "\n::int. even n \ x = of_int n * (pi / 2)" apply (simp add: sin_zero_iff, safe) - apply (metis even_int_iff real_of_int_of_nat_eq) + apply (metis even_int_iff of_int_of_nat_eq) apply (rule_tac x="- (int n)" in exI, simp) done next fix n::int assume "even n" - then show "sin (real n * (pi / 2)) = 0" + then show "sin (of_int n * (pi / 2)) = 0" apply (simp add: sin_zero_iff) apply (case_tac n rule: int_cases2, simp) apply (rule disjI2) @@ -3677,13 +3660,11 @@ qed lemma sin_zero_iff_int2: - "sin x = 0 \ (\n::int. x = real n * pi)" + "sin x = 0 \ (\n::int. x = of_int n * pi)" apply (simp only: sin_zero_iff_int) apply (safe elim!: evenE) apply (simp_all add: field_simps) - apply (subst real_numeral(1) [symmetric]) - apply (simp only: real_of_int_mult [symmetric] real_of_int_inject) - apply auto + using dvd_triv_left apply fastforce done lemma cos_monotone_0_pi: @@ -3691,7 +3672,6 @@ shows "cos x < cos y" proof - have "- (x - y) < 0" using assms by auto - from MVT2[OF \y < x\ DERIV_cos[THEN impI, THEN allI]] obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto @@ -3790,8 +3770,7 @@ have "sin ((real n + 1/2) * pi) = cos (real n * pi)" by (auto simp: algebra_simps sin_add) thus ?thesis - by (simp add: real_of_nat_Suc distrib_right add_divide_distrib - mult.commute [of pi]) + by (simp add: distrib_right add_divide_distrib add.commute mult.commute [of pi]) qed lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1" @@ -3811,7 +3790,7 @@ done lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" -by (simp only: cos_add sin_add real_of_nat_Suc distrib_right distrib_left add_divide_distrib, auto) +by (simp only: cos_add sin_add of_nat_Suc distrib_right distrib_left add_divide_distrib, auto) lemma DERIV_cos_add [simp]: "DERIV (\x. cos (x + k)) xa :> - sin (xa + k)" by (auto intro!: derivative_eq_intros) @@ -3832,9 +3811,9 @@ by simp lemma sin_times_pi_eq_0: "sin(x * pi) = 0 \ x \ \" - by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_real_of_int real_of_int_def) - -lemma cos_one_2pi: + by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_of_int) + +lemma cos_one_2pi: "cos(x) = 1 \ (\n::nat. x = n * 2*pi) | (\n::nat. x = -(n * 2*pi))" (is "?lhs = ?rhs") proof @@ -3852,7 +3831,7 @@ show ?rhs using m me n by (auto simp: field_simps elim!: evenE) - next + next fix n::nat assume n: "even n" "x = - (real n * (pi/2))" then obtain m where m: "n = 2 * m" @@ -3872,9 +3851,9 @@ lemma cos_one_2pi_int: "cos(x) = 1 \ (\n::int. x = n * 2*pi)" apply auto --\FIXME simproc bug\ apply (auto simp: cos_one_2pi) - apply (metis real_of_int_of_nat_eq) - apply (metis mult_minus_right real_of_int_minus real_of_int_of_nat_eq) - by (metis mult_minus_right of_int_of_nat real_of_int_def real_of_nat_def) + apply (metis of_int_of_nat_eq) + apply (metis mult_minus_right of_int_minus of_int_of_nat_eq) + by (metis mult_minus_right of_int_of_nat ) lemma sin_cos_sqrt: "0 \ sin(x) \ (sin(x) = sqrt(1 - (cos(x) ^ 2)))" using sin_squared_eq real_sqrt_unique by fastforce @@ -3882,7 +3861,7 @@ lemma sin_eq_0_pi: "-pi < x \ x < pi \ sin(x) = 0 \ x = 0" by (metis sin_gt_zero sin_minus minus_less_iff neg_0_less_iff_less not_less_iff_gr_or_eq) -lemma cos_treble_cos: +lemma cos_treble_cos: fixes x :: "'a::{real_normed_field,banach}" shows "cos(3 * x) = 4 * cos(x) ^ 3 - 3 * cos x" proof - @@ -3948,16 +3927,16 @@ by (simp add: sin_cos_eq cos_60) lemma cos_integer_2pi: "n \ \ \ cos(2*pi * n) = 1" - by (metis Ints_cases cos_one_2pi_int mult.assoc mult.commute real_of_int_def) + by (metis Ints_cases cos_one_2pi_int mult.assoc mult.commute) lemma sin_integer_2pi: "n \ \ \ sin(2*pi * n) = 0" by (metis sin_two_pi Ints_mult mult.assoc mult.commute sin_times_pi_eq_0) -lemma cos_int_2npi [simp]: "cos (2 * real (n::int) * pi) = 1" +lemma cos_int_2npi [simp]: "cos (2 * of_int (n::int) * pi) = 1" by (simp add: cos_one_2pi_int) -lemma sin_int_2npi [simp]: "sin (2 * real (n::int) * pi) = 0" - by (metis Ints_real_of_int mult.assoc mult.commute sin_integer_2pi) +lemma sin_int_2npi [simp]: "sin (2 * of_int (n::int) * pi) = 0" + by (metis Ints_of_int mult.assoc mult.commute sin_integer_2pi) lemma sincos_principal_value: "\y. (-pi < y \ y \ pi) \ (sin(y) = sin(x) \ cos(y) = cos(x))" apply (rule exI [where x="pi - (2*pi) * frac((pi - x) / (2*pi))"]) @@ -4201,27 +4180,29 @@ next case (Suc n) have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" - unfolding Suc_eq_plus1 real_of_nat_add real_of_one distrib_right by auto + unfolding Suc_eq_plus1 of_nat_add distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed -lemma tan_periodic_int[simp]: fixes i :: int shows "tan (x + real i * pi) = tan x" +lemma tan_periodic_int[simp]: fixes i :: int shows "tan (x + of_int i * pi) = tan x" proof (cases "0 \ i") case True - hence i_nat: "real i = real (nat i)" by auto - show ?thesis unfolding i_nat by auto + hence i_nat: "of_int i = of_int (nat i)" by auto + show ?thesis unfolding i_nat + by (metis of_int_of_nat_eq tan_periodic_nat) next case False - hence i_nat: "real i = - real (nat (-i))" by auto - have "tan x = tan (x + real i * pi - real i * pi)" + hence i_nat: "of_int i = - of_int (nat (-i))" by auto + have "tan x = tan (x + of_int i * pi - of_int i * pi)" by auto - also have "\ = tan (x + real i * pi)" - unfolding i_nat mult_minus_left diff_minus_eq_add by (rule tan_periodic_nat) + also have "\ = tan (x + of_int i * pi)" + unfolding i_nat mult_minus_left diff_minus_eq_add + by (metis of_int_of_nat_eq tan_periodic_nat) finally show ?thesis by auto qed lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x" - using tan_periodic_int[of _ "numeral n" ] unfolding real_numeral . + using tan_periodic_int[of _ "numeral n" ] by simp lemma tan_minus_45: "tan (-(pi/4)) = -1" unfolding tan_def by (simp add: sin_45 cos_45) @@ -4292,7 +4273,7 @@ lemma cot_periodic [simp]: "cot (x + 2*pi) = cot x" by (simp add: cot_def) - + lemma cot_altdef: "cot x = inverse (tan x)" by (simp add: cot_def tan_def) @@ -4494,7 +4475,7 @@ by (metis arccos_cos cos_pi order_refl pi_ge_zero) lemma arccos_minus: "-1 \ x \ x \ 1 \ arccos(-x) = pi - arccos x" - by (metis arccos_cos arccos_cos2 cos_minus_pi cos_total diff_le_0_iff_le le_add_same_cancel1 + by (metis arccos_cos arccos_cos2 cos_minus_pi cos_total diff_le_0_iff_le le_add_same_cancel1 minus_diff_eq uminus_add_conv_diff) lemma sin_arccos_nonzero: "-1 < x \ x < 1 \ ~(sin(arccos x) = 0)" @@ -4746,7 +4727,7 @@ \ (sin(x) \ sin(y) \ x \ y)" by (meson leD le_less_linear sin_monotone_2pi sin_monotone_2pi_le) -lemma sin_inj_pi: +lemma sin_inj_pi: "\-(pi/2) \ x; x \ pi/2;-(pi/2) \ y; y \ pi/2; sin(x) = sin(y)\ \ x = y" by (metis arcsin_sin) @@ -4772,24 +4753,24 @@ proof - have x1: "x \ 1" using assms - by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2) + by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2) moreover with assms have ax: "0 \ arccos x" "cos(arccos x) = x" by (auto simp: arccos) moreover have "y = sqrt (1 - x\<^sup>2)" using assms by (metis abs_of_nonneg add.commute add_diff_cancel real_sqrt_abs) - ultimately show ?thesis using assms arccos_le_pi2 [of x] + ultimately show ?thesis using assms arccos_le_pi2 [of x] by (rule_tac x="arccos x" in exI) (auto simp: sin_arccos) -qed +qed lemma sincos_total_pi: assumes "0 \ y" and "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 x]) - case le from sincos_total_pi_half [OF le] + case le from sincos_total_pi_half [OF le] show ?thesis by (metis pi_ge_two pi_half_le_two add.commute add_le_cancel_left add_mono assms) next - case ge + case ge then have "0 \ -x" by simp then obtain t where "t\0" "t \ pi/2" "-x = cos t" "y = sin t" @@ -4798,17 +4779,17 @@ by (metis \0 \ - x\ power2_minus) then show ?thesis by (rule_tac x="pi-t" in exI, auto) -qed - +qed + lemma sincos_total_2pi_le: assumes "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ 2*pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 y]) - case le from sincos_total_pi [OF le] + case le from sincos_total_pi [OF le] show ?thesis by (metis assms le_add_same_cancel1 mult.commute mult_2_right order.trans) next - case ge + case ge then have "0 \ -y" by simp then obtain t where "t\0" "t \ pi" "x = cos t" "-y = sin t" @@ -4817,7 +4798,7 @@ by (metis \0 \ - y\ power2_minus) then show ?thesis by (rule_tac x="2*pi-t" in exI, auto) -qed +qed lemma sincos_total_2pi: assumes "x\<^sup>2 + y\<^sup>2 = 1" @@ -4855,7 +4836,7 @@ done lemma arccos_le_mono: "abs x \ 1 \ abs y \ 1 \ arccos x \ arccos y \ y \ x" - using arccos_less_mono [of y x] + using arccos_less_mono [of y x] by (simp add: not_le [symmetric]) lemma arccos_less_arccos: "-1 \ x \ x < y \ y \ 1 \ arccos y < arccos x" @@ -4934,7 +4915,7 @@ moreover have "\336/527\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF 17] this] - have "arctan(1/7) + arctan (336/527) = arctan (2879/3353)" by auto + have "arctan(1/7) + arctan (336/527) = arctan (2879/3353)" by auto ultimately have I: "5 * arctan(1/7) = arctan (2879/3353)" by auto have 379: "\3/79\ < (1 :: real)" by auto with arctan_double @@ -5272,7 +5253,7 @@ } have "?a 1 ----> 0" unfolding tendsto_rabs_zero_iff power_one divide_inverse One_nat_def - by (auto intro!: tendsto_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat simp del: real_of_nat_Suc) + by (auto intro!: tendsto_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat simp del: of_nat_Suc) have "?diff 1 ----> 0" proof (rule LIMSEQ_I) fix r :: real @@ -5459,7 +5440,7 @@ lemma polynomial_product: (*with thanks to Chaitanya Mangla*) fixes x:: "'a :: idom" assumes m: "\i. i>m \ (a i) = 0" and n: "\j. j>n \ (b j) = 0" - shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = + shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" proof - have "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\i\m. \j\n. (a i * x ^ i) * (b j * x ^ j))" @@ -5473,19 +5454,19 @@ also have "... = (\(i,j)\{(i,j). i+j \ m+n}. (a i * x ^ i) * (b j * x ^ j))" by (auto simp: pairs_le_eq_Sigma setsum.Sigma) also have "... = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" - apply (subst setsum_triangle_reindex_eq) + apply (subst setsum_triangle_reindex_eq) apply (auto simp: algebra_simps setsum_right_distrib intro!: setsum.cong) by (metis le_add_diff_inverse power_add) finally show ?thesis . qed -lemma polynomial_product_nat: +lemma polynomial_product_nat: fixes x:: nat assumes m: "\i. i>m \ (a i) = 0" and n: "\j. j>n \ (b j) = 0" - shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = + shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" using polynomial_product [of m a n b x] assms - by (simp add: Int.zpower_int Int.zmult_int Int.int_setsum [symmetric]) + by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric] int_int_eq Int.int_setsum [symmetric]) lemma polyfun_diff: (*COMPLEX_SUB_POLYFUN in HOL Light*) fixes x :: "'a::idom" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/ex/Ballot.thy --- a/src/HOL/ex/Ballot.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/ex/Ballot.thy Tue Nov 10 14:18:41 2015 +0000 @@ -266,7 +266,7 @@ proof - from main_nat[of a b] \b < a\ have "(real a + real b) * real (valid_countings a b) = (real a - real b) * real (all_countings a b)" - by (simp only: real_of_nat_add[symmetric] real_of_nat_mult[symmetric]) auto + by (simp only: of_nat_add[symmetric] of_nat_mult[symmetric]) auto from this \b < a\ show ?thesis by (subst mult_left_cancel[of "real a + real b", symmetric]) auto qed diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/ex/HarmonicSeries.thy --- a/src/HOL/ex/HarmonicSeries.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/ex/HarmonicSeries.thy Tue Nov 10 14:18:41 2015 +0000 @@ -72,9 +72,7 @@ thus ?thesis by simp qed moreover from xs have "x \ 2^m" by auto - ultimately have - "inverse (real x) \ inverse (real ((2::nat)^m))" - by (simp del: real_of_nat_power) + ultimately have "inverse (real x) \ inverse (real ((2::nat)^m))" by simp moreover from xgt0 have "real x \ 0" by simp then have @@ -99,7 +97,7 @@ by simp also have "\ = ((1/(real tm)) * real (card (?S m)))" - by (simp add: real_of_card real_of_nat_def) + by (simp add: real_of_card) also have "\ = ((1/(real tm)) * real (tm - (2^(m - 1))))" by (simp add: tmdef) @@ -280,28 +278,13 @@ let ?s = "suminf ?f" -- "let ?s equal the sum of the harmonic series" assume sf: "summable ?f" then obtain n::nat where ndef: "n = nat \2 * ?s\" by simp - then have ngt: "1 + real n/2 > ?s" - proof - - have "\n. 0 \ ?f n" by simp - with sf have "?s \ 0" - by (rule suminf_nonneg) - then have cgt0: "\2*?s\ \ 0" by simp - - from ndef have "n = nat \(2*?s)\" . - then have "real n = real (nat \2*?s\)" by simp - with cgt0 have "real n = real \2*?s\" - by (auto dest: real_nat_eq_real) - then have "real n \ 2*(?s)" by simp - then have "real n/2 \ (?s)" by simp - then show "1 + real n/2 > (?s)" by simp - qed - - obtain j where jdef: "j = (2::nat)^n" by simp + then have ngt: "1 + real n/2 > ?s" by linarith + def j \ "(2::nat)^n" have "\m\j. 0 < ?f m" by simp with sf have "(\ii\{Suc 0..i\{1..< Suc ((2::nat)^n)}. 1 / (real i)) < ?s" by simp then have "(\i\{1..(2::nat)^n}. 1 / (real i)) < ?s" diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/ex/Sqrt_Script.thy --- a/src/HOL/ex/Sqrt_Script.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/ex/Sqrt_Script.thy Tue Nov 10 14:18:41 2015 +0000 @@ -60,8 +60,8 @@ "prime (p::nat) \ x * x = real p \ 0 \ x \ x \ \" apply (rule notI) apply (erule Rats_abs_nat_div_natE) - apply (simp del: real_of_nat_mult - add: abs_if divide_eq_eq prime_not_square real_of_nat_mult [symmetric]) + apply (simp del: of_nat_mult + add: abs_if divide_eq_eq prime_not_square of_nat_mult [symmetric]) done lemmas two_sqrt_irrational = diff -r 933eb9e6a1cc -r 77b453bd616f src/HOL/ex/Sum_of_Powers.thy --- a/src/HOL/ex/Sum_of_Powers.thy Tue Nov 03 11:20:21 2015 +0100 +++ b/src/HOL/ex/Sum_of_Powers.thy Tue Nov 10 14:18:41 2015 +0000 @@ -1,7 +1,4 @@ -(* Title: HOL/ex/Sum_of_Powers.thy - Author: Lukas Bulwahn -*) - +(* Author: Lukas Bulwahn *) section \Sum of Powers\ theory Sum_of_Powers @@ -32,8 +29,8 @@ by (simp only: Suc_eq_plus1) finally show ?thesis . qed - from this have "(\i 'a)) (n + 1 - k) / of_nat (n + 1) * (\ii 'a)) (n + 1 - k) / of_nat (n + 1) * (\i n" shows "(n choose k) = (n + 1 - k) / (n + 1) * (Suc n choose k)" -proof - - have "real (n choose k) = of_nat (n choose k)" by auto - also have "... = of_nat (n + 1 - k) / of_nat (n + 1) * of_nat (Suc n choose k)" - by (simp add: assms of_nat_binomial_eq_mult_binomial_Suc) - also have "... = (n + 1 - k) / (n + 1) * (Suc n choose k)" - using real_of_nat_def by auto - finally show ?thesis - by (metis (no_types, lifting) assms le_add1 le_trans of_nat_diff real_of_nat_1 real_of_nat_add real_of_nat_def) -qed +by (metis Suc_eq_plus1 add.commute assms le_SucI of_nat_Suc of_nat_binomial_eq_mult_binomial_Suc of_nat_diff) subsection \Preliminaries\ @@ -90,7 +79,7 @@ lemma bernpoly_0: "bernpoly n 0 = bernoulli n" proof (cases n) case 0 - from this show "bernpoly n 0 = bernoulli n" + then show "bernpoly n 0 = bernoulli n" unfolding bernpoly_def bernoulli.simps by auto next case (Suc n') @@ -104,10 +93,10 @@ "(\k\n. ((Suc n) choose k) * bernoulli k) = (if n = 0 then 1 else 0)" proof (cases n) case 0 - from this show ?thesis by (simp add: bernoulli.simps) + then show ?thesis by (simp add: bernoulli.simps) next case Suc - from this show ?thesis + then show ?thesis by (simp add: bernoulli.simps) (simp add: field_simps add_2_eq_Suc'[symmetric] del: add_2_eq_Suc add_2_eq_Suc') qed @@ -121,7 +110,7 @@ unfolding bernpoly_def by (rule DERIV_cong) (fast intro!: derivative_intros, simp) moreover have "(\k\n. real (Suc n - k) * x ^ (n - k) * (real (Suc n choose k) * bernoulli k)) = (n + 1) * bernpoly n x" unfolding bernpoly_def - by (auto intro: setsum.cong simp add: setsum_right_distrib real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 real_of_nat_diff) + by (auto intro: setsum.cong simp add: setsum_right_distrib real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 of_nat_diff) ultimately show ?thesis by auto qed @@ -134,7 +123,7 @@ case (Suc n) have "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = (Suc n) * 0 ^ n" unfolding bernpoly_0 unfolding bernpoly_def by (simp add: setsum_binomial_times_bernoulli zero_power) - from this have const: "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = real (Suc n) * 0 ^ n" by (simp add: power_0_left) + then have const: "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = real (Suc n) * 0 ^ n" by (simp add: power_0_left) have hyps': "\x. (real n + 1) * bernpoly n (x + 1) - (real n + 1) * bernpoly n x = real n * x ^ (n - Suc 0) * real (Suc n)" unfolding right_diff_distrib[symmetric] by (simp add: Suc.hyps One_nat_def) note [derivative_intros] = DERIV_chain'[where f = "\x::real. x + 1" and g = "bernpoly (Suc n)" and s="UNIV"] @@ -148,7 +137,7 @@ from diff_bernpoly[of "Suc m", simplified] have "(m + (1::real)) * (\k\n. (real k) ^ m) = (\k\n. bernpoly (Suc m) (real k + 1) - bernpoly (Suc m) (real k))" by (auto simp add: setsum_right_distrib intro!: setsum.cong) also have "... = (\k\n. bernpoly (Suc m) (real (k + 1)) - bernpoly (Suc m) (real k))" - by (simp only: real_of_nat_1[symmetric] real_of_nat_add[symmetric]) + by simp also have "... = bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0" by (simp only: setsum_diff[where f="\k. bernpoly (Suc m) (real k)"]) simp finally show ?thesis by (auto simp add: field_simps intro!: eq_divide_imp) @@ -185,9 +174,9 @@ proof - from sum_of_squares have "real (6 * (\k\n. k ^ 2)) = real (2 * n ^ 3 + 3 * n ^ 2 + n)" by (auto simp add: field_simps) - from this have "6 * (\k\n. k ^ 2) = 2 * n ^ 3 + 3 * n ^ 2 + n" - by (simp only: real_of_nat_inject[symmetric]) - from this show ?thesis by auto + then have "6 * (\k\n. k ^ 2) = 2 * n ^ 3 + 3 * n ^ 2 + n" + by blast + then show ?thesis by auto qed lemma sum_of_cubes: "(\k\n::nat. k ^ 3) = (n ^ 2 + n) ^ 2 / 4" @@ -202,14 +191,14 @@ by (simp add: unroll algebra_simps power2_eq_square power4_eq power3_eq_cube) finally show ?thesis by simp qed - + lemma sum_of_cubes_nat: "(\k\n::nat. k ^ 3) = (n ^ 2 + n) ^ 2 div 4" proof - from sum_of_cubes have "real (4 * (\k\n. k ^ 3)) = real ((n ^ 2 + n) ^ 2)" by (auto simp add: field_simps) - from this have "4 * (\k\n. k ^ 3) = (n ^ 2 + n) ^ 2" - by (simp only: real_of_nat_inject[symmetric]) - from this show ?thesis by auto + then have "4 * (\k\n. k ^ 3) = (n ^ 2 + n) ^ 2" + by blast + then show ?thesis by auto qed end