diff -r 54a3db2ed201 -r 903bb1495239 src/HOL/Library/Float.thy --- a/src/HOL/Library/Float.thy Wed Jun 17 10:57:11 2015 +0200 +++ b/src/HOL/Library/Float.thy Wed Jun 17 11:03:05 2015 +0200 @@ -3,7 +3,7 @@ Copyright 2012 TU München *) -section {* Floating-Point Numbers *} +section \Floating-Point Numbers\ theory Float imports Complex_Main Lattice_Algebras @@ -43,7 +43,7 @@ lemma real_float[simp]: "x \ float \ real (float_of x) = x" unfolding real_of_float_def by (rule float_of_inverse) -subsection {* Real operations preserving the representation as floating point number *} +subsection \Real operations preserving the representation as floating point number\ lemma floatI: fixes m e :: int shows "m * 2 powr e = x \ x \ float" by (auto simp: float_def) @@ -161,7 +161,7 @@ code_datatype Float -subsection {* Arithmetic operations on floating point numbers *} +subsection \Arithmetic operations on floating point numbers\ instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}" begin @@ -264,7 +264,7 @@ and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)" unfolding real_of_float_eq by simp_all -subsection {* Quickcheck *} +subsection \Quickcheck\ instantiation float :: exhaustive begin @@ -304,7 +304,7 @@ end -subsection {* Represent floats as unique mantissa and exponent *} +subsection \Represent floats as unique mantissa and exponent\ lemma int_induct_abs[case_names less]: fixes j :: int @@ -320,7 +320,7 @@ case (less n) { fix m assume n: "n \ 0" "n = m * r" then have "\m \ < \n\" - using `1 < r` by (simp add: abs_mult) + using \1 < r\ by (simp add: abs_mult) from less[OF this] n have "\k i. n = k * r ^ Suc i \ \ r dvd k" by auto } then show ?case by (metis dvd_def monoid_mult_class.mult.right_neutral mult.commute power_0) @@ -333,7 +333,7 @@ proof have "m1 \ 0" using m1 unfolding dvd_def by auto assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2" - with `e1 \ e2` have "m1 = m2 * 2 powr nat (e2 - e1)" + with \e1 \ e2\ have "m1 = m2 * 2 powr nat (e2 - e1)" by (simp add: powr_divide2[symmetric] field_simps) also have "\ = m2 * 2^nat (e2 - e1)" by (simp add: powr_realpow) @@ -342,7 +342,7 @@ with m1 have "m1 = m2" by (cases "nat (e2 - e1)") (auto simp add: dvd_def) then show "m1 = m2 \ e1 = e2" - using eq `m1 \ 0` by (simp add: powr_inj) + using eq \m1 \ 0\ by (simp add: powr_inj) qed simp lemma mult_powr_eq_mult_powr_iff: @@ -359,9 +359,9 @@ proof atomize_elim { assume "x \ 0" from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def) - with `x \ 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\ 2 dvd k" + with \x \ 0\ int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\ 2 dvd k" by auto - with `\ 2 dvd k` x have "\(m::int) (e::int). x = m * 2 powr e \ \ (2::int) dvd m" + with \\ 2 dvd k\ x have "\(m::int) (e::int). x = m * 2 powr e \ \ (2::int) dvd m" by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"]) (simp add: powr_add powr_realpow) } then show "x = 0 \ (\(m::int) (e::int). x = m * 2 powr e \ \ (2::int) dvd m \ x \ 0)" @@ -434,7 +434,7 @@ by (auto simp: mult_powr_eq_mult_powr_iff) qed -subsection {* Compute arithmetic operations *} +subsection \Compute arithmetic operations\ lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f" unfolding real_of_float_eq mantissa_exponent[of f] by simp @@ -467,7 +467,7 @@ using eq by simp then have "mantissa f = m * 2^nat (e - exponent f)" unfolding real_of_int_inject by simp - with `exponent f < e` have "2 dvd mantissa f" + 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)") apply auto @@ -476,7 +476,7 @@ qed ultimately have "real m = mantissa f * 2^nat (exponent f - e)" by (simp add: powr_realpow[symmetric]) - with `e \ exponent f` + 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 qed @@ -564,7 +564,7 @@ hide_fact (open) compute_float_eq -subsection {* Lemmas for types @{typ real}, @{typ nat}, @{typ int}*} +subsection \Lemmas for types @{typ real}, @{typ nat}, @{typ int}\ lemmas real_of_ints = real_of_int_zero @@ -588,7 +588,7 @@ lemmas nat_of_reals = real_of_nats[symmetric] -subsection {* Rounding Real Numbers *} +subsection \Rounding Real Numbers\ definition round_down :: "int \ real \ real" where "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec" @@ -663,9 +663,9 @@ proof - have "x * 2 powr p < 1 / 2 * 2 powr p" using assms by simp - also have "\ \ 2 powr p - 1" using `p > 0` + 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` + finally show ?thesis using \p > 0\ by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_eq) qed @@ -705,7 +705,7 @@ by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff) -subsection {* Rounding Floats *} +subsection \Rounding Floats\ definition div_twopow::"int \ nat \ int" where [simp]: "div_twopow x n = x div (2 ^ n)" @@ -763,7 +763,7 @@ also have "... = 1 / 2 powr p / 2 powr e" unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add) finally show ?thesis - using `p + e < 0` + using \p + e < 0\ by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric]) next assume "\ p + e < 0" @@ -771,7 +771,7 @@ have r: "\(m * 2 powr e) * 2 powr real p\ = (m * 2 powr e) * 2 powr real 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 + with \\ p + e < 0\ show ?thesis by transfer (auto simp add: round_down_def field_simps powr_add powr_minus) qed hide_fact (open) compute_float_down @@ -791,16 +791,16 @@ proof cases assume "\ b dvd a" hence "a mod b \ 0" by auto - hence ne: "real (a mod b) / real b \ 0" using `b \ 0` by auto + hence ne: "real (a mod b) / real b \ 0" using \b \ 0\ by auto have "\real a / real b\ = \real a / real b\ + 1" apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric]) proof - have "real \real a / real b\ \ real a / real b" by simp moreover have "real \real a / real b\ \ real a / real b" - apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \ 0` by auto + apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne \b \ 0\ by auto ultimately show "real \real a / real b\ < real a / real b" by arith qed - thus ?thesis using `\ b dvd a` by simp + thus ?thesis using \\ b dvd a\ by simp qed (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) @@ -810,7 +810,7 @@ hide_fact (open) compute_float_up -subsection {* Compute bitlen of integers *} +subsection \Compute bitlen of integers\ definition bitlen :: "int \ int" where "bitlen a = (if a > 0 then \log 2 a\ + 1 else 0)" @@ -820,7 +820,7 @@ { assume "0 > x" have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all - also have "... < log 2 (-x)" using `0 > x` by auto + also have "... < log 2 (-x)" using \0 > x\ by auto finally have "-1 < log 2 (-x)" . } thus "0 \ bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg) qed @@ -830,22 +830,22 @@ shows "2 ^ nat (bitlen x - 1) \ x \ x < 2 ^ nat (bitlen 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 powr_realpow[symmetric, of 2 "nat \log 2 (real x)\"] \x > 0\ using real_nat_eq_real[of "floor (log 2 (real x))"] by simp also have "... \ 2 powr log 2 (real x)" by simp also have "... = real x" - using `0 < x` by simp + using \0 < x\ by simp finally have "2 ^ nat \log 2 (real x)\ \ real x" by simp - thus "2 ^ nat (bitlen x - 1) \ x" using `x > 0` + thus "2 ^ nat (bitlen x - 1) \ x" using \x > 0\ by (simp add: bitlen_def) next - have "x \ 2 powr (log 2 x)" using `x > 0` by simp + have "x \ 2 powr (log 2 x)" using \x > 0\ by simp also have "... < 2 ^ nat (\log 2 (real x)\ + 1)" apply (simp add: powr_realpow[symmetric]) - using `x > 0` by simp - finally show "x < 2 ^ nat (bitlen x)" using `x > 0` + using \x > 0\ by simp + finally show "x < 2 ^ nat (bitlen x)" using \x > 0\ by (simp add: bitlen_def ac_simps) qed @@ -874,7 +874,7 @@ by (simp add: mantissa_noteq_0) moreover obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i" - by (rule f_def[THEN denormalize_shift, OF `f \ float_of 0`]) + by (rule f_def[THEN denormalize_shift, OF \f \ float_of 0\]) ultimately show ?thesis by (simp add: abs_mult) qed @@ -890,28 +890,28 @@ next def n \ "\log 2 (real x)\" then have "0 \ n" - using `2 \ x` by simp + using \2 \ x\ by simp assume "x mod 2 \ 0" - with `2 \ x` have "x mod 2 = 1" "\ 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0) - with `2 \ x` have "x \ 2^nat n" by (cases "nat n") auto + with \2 \ x\ have "x mod 2 = 1" "\ 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0) + with \2 \ x\ have "x \ 2^nat n" by (cases "nat n") auto moreover { have "real (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 } + 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 { have "n = \log 2 (2^nat n)\" - using `0 \ n` by (simp add: log_nat_power) + 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 (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) + using \2 \ x\ by (auto simp add: n_def intro!: floor_mono) ultimately show "\log 2 (x - x mod 2)\ = \log 2 x\" - unfolding n_def `x mod 2 = 1` by auto + unfolding n_def \x mod 2 = 1\ by auto qed finally have "\log 2 (x div 2)\ + 1 = \log 2 x\" . } moreover @@ -934,7 +934,7 @@ hence "m \ 0" by auto show ?thesis proof (cases "0 \ e") - case True thus ?thesis using `0 < m` by (simp add: bitlen_def) + case True thus ?thesis using \0 < m\ by (simp add: bitlen_def) next have "(1::int) < 2" by simp case False let ?S = "2^(nat (-e))" @@ -945,8 +945,8 @@ hence "1 * ?S \ real m * inverse ?S * ?S" by (rule mult_right_mono, auto) hence "?S \ real m" unfolding mult.assoc by auto hence "?S \ m" unfolding real_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] + from this bitlen_bounds[OF \0 < m\, THEN conjunct2] + have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF \1 < 2\, symmetric] by (rule order_le_less_trans) hence "-e < bitlen m" using False by auto thus ?thesis by auto @@ -959,22 +959,22 @@ proof - let ?B = "2^nat(bitlen m - 1)" - have "?B \ m" using bitlen_bounds[OF `0 m" using bitlen_bounds[OF \0 ] .. hence "1 * ?B \ real m" unfolding real_of_int_le_iff[symmetric] by auto thus "1 \ real m / ?B" by auto have "m \ 0" using assms by auto - have "0 \ bitlen m - 1" using `0 < m` by (auto simp: bitlen_def) + have "0 \ bitlen m - 1" using \0 < m\ by (auto simp: bitlen_def) - have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 = 2^nat(bitlen m - 1 + 1)" 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 + have "m < 2^nat(bitlen m)" using bitlen_bounds[OF \0 ] .. + also have "\ = 2^nat(bitlen m - 1 + 1)" 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 hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto) thus "real m / ?B < 2" by auto qed -subsection {* Truncating Real Numbers*} +subsection \Truncating Real Numbers\ definition truncate_down::"nat \ real \ real" where "truncate_down prec x = round_down (prec - \log 2 \x\\ - 1) x" @@ -1051,7 +1051,7 @@ } ultimately show ?thesis by arith qed -subsection {* Truncating Floats*} +subsection \Truncating Floats\ lift_definition float_round_up :: "nat \ float \ float" is truncate_up by (simp add: truncate_up_def) @@ -1093,7 +1093,7 @@ hide_fact (open) compute_float_round_up -subsection {* Approximation of positive rationals *} +subsection \Approximation of positive rationals\ lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)" by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps) @@ -1146,21 +1146,21 @@ 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'" - by (simp add: powr_realpow[symmetric] `0 \ l` x'_def) + 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` + 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) next assume "\ 0 \ l" def y' \ "y * 2 ^ nat (- l)" - from `y \ 0` have "y' \ 0" by (simp add: y'_def) + 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'" - using `\ 0 \ l` + 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` + 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]) @@ -1214,7 +1214,7 @@ by transfer (simp add: round_down_uminus_eq) hide_fact (open) compute_rapprox_rat -subsection {* Division *} +subsection \Division\ definition "real_divl prec a b = round_down (int prec + \ log 2 \b\ \ - \ log 2 \a\ \) (a / b)" @@ -1250,7 +1250,7 @@ hide_fact (open) compute_float_divr -subsection {* Approximate Power *} +subsection \Approximate Power\ lemma div2_less_self[termination_simp]: fixes n::nat shows "odd n \ n div 2 < n" by (simp add: odd_pos) @@ -1306,9 +1306,9 @@ also have "\ = x ^ (Suc n div 2 * 2)" by (simp add: power_mult[symmetric]) also have "Suc n div 2 * 2 = Suc n" - using `odd n` by presburger + using \odd n\ by presburger finally have ?case - using `odd n` + using \odd n\ by (auto intro!: truncate_down_le simp del: odd_Suc_div_two) } thus ?case by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg) @@ -1320,14 +1320,14 @@ { assume "odd n" hence "Suc n = Suc n div 2 * 2" - using `odd n` even_Suc by presburger + using \odd n\ even_Suc by presburger hence "x ^ Suc n \ (x ^ (Suc n div 2))\<^sup>2" by (simp add: power_mult[symmetric]) also have "\ \ (power_up p x (Suc n div 2))\<^sup>2" - using 2 `odd n` + using 2 \odd n\ by (auto intro: power_mono simp del: odd_Suc_div_two ) finally have ?case - using `odd n` + using \odd n\ by (auto intro!: truncate_up_le simp del: odd_Suc_div_two ) } thus ?case by (auto intro!: truncate_up_le mult_left_mono 2) @@ -1350,7 +1350,7 @@ by transfer simp -subsection {* Approximate Addition *} +subsection \Approximate Addition\ definition "plus_down prec x y = truncate_down prec (x + y)" @@ -1432,7 +1432,7 @@ also note b_le_1 finally have b_less_1: "b * 2 powr real 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 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)\" @@ -1474,12 +1474,12 @@ also have "\ = \(2 * ai + b * 2 powr (p + 1)) / real ((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))\" - using `b < 0` assms + 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) also have "\ = \(2 * ai - 1) * 2 powr (q - p - 1)\" using assms by (simp add: algebra_simps divide_powr_uminus powr_realpow[symmetric]) - finally have ?thesis using `b < 0` by simp + finally have ?thesis using \b < 0\ by simp } ultimately show ?thesis by arith qed @@ -1495,37 +1495,37 @@ def k \ "\log 2 \ai\\" hence "\log 2 \ai\\ = k" by simp hence k: "2 powr k \ \ai\" "\ai\ < 2 powr (k + 1)" - by (simp_all add: floor_log_eq_powr_iff `ai \ 0`) + by (simp_all add: floor_log_eq_powr_iff \ai \ 0\) have "k \ 0" using assms by (auto simp: k_def) def r \ "\ai\ - 2 ^ nat k" have r: "0 \ r" "r < 2 powr k" - using `k \ 0` k + using \k \ 0\ k by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int) hence "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` + using \k \ 0\ by (auto simp: powr_int) + from this[simplified real_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) have "\ai\ = 2 powr k + r" - using `k \ 0` by (auto simp: k_def r_def powr_realpow[symmetric]) + using \k \ 0\ by (auto simp: k_def r_def powr_realpow[symmetric]) have pos: "\b::real. abs b < 1 \ 0 < 2 powr k + (r + b)" - using `0 \ k` `ai \ 0` + using \0 \ k\ \ai \ 0\ by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps split: split_if_asm) have less: "\sgn ai * b\ < 1" and less': "\sgn (sgn ai * b) / 2\ < 1" - using `abs b \ _` by (auto simp: abs_if sgn_if split: split_if_asm) + using \abs b \ _\ by (auto simp: abs_if sgn_if split: split_if_asm) have floor_eq: "\b::real. abs b \ 1 / 2 \ \log 2 (1 + (r + b) / 2 powr k)\ = (if r = 0 \ b < 0 then -1 else 0)" - using `k \ 0` r r_le + 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)" - using `abs b <= _` `0 \ k` r + from \real \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)\" proof - @@ -1537,14 +1537,14 @@ also let ?if = "if r = 0 \ sgn ai * b < 0 then -1 else 0" have "\log 2 (1 + (r + sgn ai * b) / 2 powr k)\ = ?if" - using `abs b <= _` + using \abs b <= _\ by (intro floor_eq) (auto simp: abs_mult sgn_if) also have "\ = \log 2 (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k)\" by (subst floor_eq) (auto simp: sgn_if) also have "k + \ = \log 2 (2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k))\" unfolding floor_add2[symmetric] - using pos[OF less'] `abs b \ _` + using pos[OF less'] \abs b \ _\ by (simp add: field_simps add_log_eq_powr) also have "2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k) = 2 powr k + r + sgn (sgn ai * b) / 2" @@ -1552,7 +1552,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 \ai\ = _\[symmetric] using \ai \ 0\ by (auto simp: abs_if sgn_if algebra_simps) finally show ?thesis . qed @@ -1590,7 +1590,7 @@ by simp finally have b_less_quarter: "\?b\ < 1/4 * 2 powr real 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 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) hence a_half_less_sum: "\?a\ / 2 < \?sum\" @@ -1600,7 +1600,7 @@ by simp_all have "\real (Float m1 e1)\ \ 1/4 * 2 powr real e1" - using `m1 \ 0` + using \m1 \ 0\ by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult) hence "?sum \ 0" using b_less_quarter by (rule sum_neq_zeroI) @@ -1608,16 +1608,16 @@ unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff) have "\real ?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) + using \m1 \ 0\ \m2 \ 0\ by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps) hence 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" - using `?m1 + ?m2 \ 0` + 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\\" - using abs_m2_less_half `m1 \ 0` + 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" by (auto simp: sgn_if zero_less_mult_iff less_not_sym) @@ -1625,7 +1625,7 @@ 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) hence "\log 2 \?m1 + ?m2'\\ + ?e = \log 2 \real (Float (?m1 * 2 + sgn m2) (?e - 1))\\" - using `?m1 + ?m2' \ 0` + using \?m1 + ?m2' \ 0\ unfolding floor_add[symmetric] by (simp add: log_add_eq_powr abs_mult_pos) finally @@ -1645,16 +1645,16 @@ 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" - using `m1 \ 0` + 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` + 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) finally have "int p - \log 2 \?a + ?b\\ \ p - (bitlen \m1\) - e1 + 2" - by (auto simp: algebra_simps bitlen_def `m1 \ 0`) + by (auto simp: algebra_simps bitlen_def \m1 \ 0\) also have "\ \ 1 - ?e" using bitlen_nonneg[of "\m1\"] by (simp add: k1_def) finally show "?f \ - ?e" by simp @@ -1707,7 +1707,7 @@ by (metis mantissa_0 zero_float.abs_eq) -subsection {* Lemmas needed by Approximate *} +subsection \Lemmas needed by Approximate\ lemma Float_num[simp]: shows "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and @@ -1804,7 +1804,7 @@ have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent) also have "mantissa x \ \mantissa x\" by simp also have "... \ 2 powr (bitlen \mantissa x\)" - using bitlen_bounds[of "\mantissa x\"] bitlen_nonneg `mantissa x \ 0` + using bitlen_bounds[of "\mantissa x\"] bitlen_nonneg \mantissa x \ 0\ by (auto simp del: real_of_int_abs simp add: powr_int) finally show ?thesis by (simp add: powr_add) qed @@ -1813,7 +1813,7 @@ 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\" using `prec \ 1` by arith + have "log 2 x \ real prec + real \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 @@ -1827,7 +1827,7 @@ lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x \ 1" shows "1 \ real_divr prec 1 x" proof - - have "1 \ 1 / x" using `0 < x` and `x <= 1` by auto + have "1 \ 1 / x" using \0 < x\ and \x <= 1\ by auto also have "\ \ real_divr prec 1 x" using real_divr[where x=1 and y=x] by auto finally show ?thesis by auto qed @@ -1877,7 +1877,7 @@ 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) thus "x * 2 powr real (int prec - \log 2 x\ - 1) \ real ((2::int) ^ prec)" - using `0 < x` by (simp add: powr_realpow) + using \0 < x\ by (simp add: powr_realpow) qed hence "real \x * 2 powr real (int prec - \log 2 x\ - 1)\ \ 2 powr int prec" by (auto simp: powr_realpow) @@ -1885,14 +1885,14 @@ have "2 powr - real (int prec - \log 2 x\ - 1) \ 2 powr - real (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\))" - using assms `0 < x` + 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\)" by simp also have "\ = 2 powr (log 2 y + real (int prec - \log 2 y\) - real (int prec - \log 2 y\))" by (subst powr_add[symmetric]) simp also have "\ = y" - using `0 < x` assms + using \0 < x\ assms by (simp add: powr_add) also have "\ \ truncate_up prec y" by (rule truncate_up) @@ -1910,8 +1910,8 @@ assumes "x \ 0" "0 \ y" shows "truncate_up prec x \ truncate_up prec y" proof - - note truncate_up_nonpos[OF `x \ 0`] - also note truncate_up_le[OF `0 \ y`] + note truncate_up_nonpos[OF \x \ 0\] + also note truncate_up_le[OF \0 \ y\] finally show ?thesis . qed @@ -1922,7 +1922,7 @@ have "x * 2 powr (- real \log 2 x\ - 1) = x * inverse (2 powr ((real \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)" - using `0 < x` + 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 @@ -1933,7 +1933,7 @@ by simp moreover have "0 \ \x * 2 powr (- real \log 2 x\ - 1)\" - using `x > 0` by auto + using \x > 0\ by auto ultimately have "\x * 2 powr (- real \log 2 x\ - 1)\ \ {0 ..< 1}" by simp also have "\ \ {0}" by auto @@ -1947,8 +1947,8 @@ assumes "x \ y" shows "truncate_down prec x \ truncate_down prec y" proof - - note truncate_down_le[OF `x \ 0`] - also note truncate_down_nonneg[OF `0 \ y`] + note truncate_down_le[OF \x \ 0\] + also note truncate_down_nonneg[OF \0 \ y\] finally show ?thesis . qed @@ -1976,33 +1976,33 @@ moreover assume "\log 2 \x\\ \ \log 2 \y\\" ultimately have logless: "log 2 x < log 2 y" and flogless: "\log 2 x\ < \log 2 y\" - unfolding atomize_conj abs_of_pos[OF `0 < x`] abs_of_pos[OF `0 < y`] + unfolding atomize_conj abs_of_pos[OF \0 < x\] abs_of_pos[OF \0 < y\] by (metis floor_less_cancel linorder_cases not_le) assume "prec \ 0" hence [simp]: "prec \ Suc 0" by auto have "2 powr (prec - 1) \ y * 2 powr real (prec - 1) / (2 powr log 2 y)" - using `0 < y` + using \0 < y\ by simp also have "\ \ y * 2 powr real prec / (2 powr (real \log 2 y\ + 1))" - using `0 \ y` `0 \ x` assms(2) + using \0 \ y\ \0 \ x\ assms(2) by (auto intro!: powr_mono divide_left_mono simp: real_of_nat_diff powr_add powr_divide2[symmetric]) also have "\ = y * 2 powr real prec / (2 powr real \log 2 y\ * 2)" by (auto simp: powr_add) finally have "(2 ^ (prec - 1)) \ \y * 2 powr real (int prec - \log 2 \y\\ - 1)\" - using `0 \ y` + using \0 \ y\ by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow) hence "(2 ^ (prec - 1)) * 2 powr - real (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 + 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)" 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\\)" - using logless flogless `x > 0` `y > 0` + 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)