src/HOL/Library/Float.thy
author Thomas Sewell <thomas.sewell@nicta.com.au>
Wed, 11 Jun 2014 14:24:23 +1000
changeset 58834 74bf65a1910a
parent 58119 9c3f0ae99532
child 58854 cc97b347b301
permissions -rw-r--r--
Hypsubst preserves equality hypotheses

Fixes included for various theories affected by this change.
     1 (*  Title:      HOL/Library/Float.thy
     2     Author:     Johannes Hölzl, Fabian Immler
     3     Copyright   2012  TU München
     4 *)
     5 
     6 header {* Floating-Point Numbers *}
     7 
     8 theory Float
     9 imports Complex_Main Lattice_Algebras
    10 begin
    11 
    12 definition "float = {m * 2 powr e | (m :: int) (e :: int). True}"
    13 
    14 typedef float = float
    15   morphisms real_of_float float_of
    16   unfolding float_def by auto
    17 
    18 defs (overloaded)
    19   real_of_float_def[code_unfold]: "real \<equiv> real_of_float"
    20 
    21 lemma type_definition_float': "type_definition real float_of float"
    22   using type_definition_float unfolding real_of_float_def .
    23 
    24 setup_lifting (no_code) type_definition_float'
    25 
    26 lemmas float_of_inject[simp]
    27 
    28 declare [[coercion "real :: float \<Rightarrow> real"]]
    29 
    30 lemma real_of_float_eq:
    31   fixes f1 f2 :: float shows "f1 = f2 \<longleftrightarrow> real f1 = real f2"
    32   unfolding real_of_float_def real_of_float_inject ..
    33 
    34 lemma float_of_real[simp]: "float_of (real x) = x"
    35   unfolding real_of_float_def by (rule real_of_float_inverse)
    36 
    37 lemma real_float[simp]: "x \<in> float \<Longrightarrow> real (float_of x) = x"
    38   unfolding real_of_float_def by (rule float_of_inverse)
    39 
    40 subsection {* Real operations preserving the representation as floating point number *}
    41 
    42 lemma floatI: fixes m e :: int shows "m * 2 powr e = x \<Longrightarrow> x \<in> float"
    43   by (auto simp: float_def)
    44 
    45 lemma zero_float[simp]: "0 \<in> float" by (auto simp: float_def)
    46 lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp
    47 lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp
    48 lemma neg_numeral_float[simp]: "- numeral i \<in> float" by (intro floatI[of "- numeral i" 0]) simp
    49 lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp
    50 lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float" by (intro floatI[of x 0]) simp
    51 lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float" by (intro floatI[of 1 i]) simp
    52 lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float" by (intro floatI[of 1 i]) simp
    53 lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float" by (intro floatI[of 1 "-i"]) simp
    54 lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" by (intro floatI[of 1 "-i"]) simp
    55 lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp
    56 lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float" by (intro floatI[of 1 "- numeral i"]) simp
    57 lemma two_pow_float[simp]: "2 ^ n \<in> float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow)
    58 lemma real_of_float_float[simp]: "real (f::float) \<in> float" by (cases f) simp
    59 
    60 lemma plus_float[simp]: "r \<in> float \<Longrightarrow> p \<in> float \<Longrightarrow> r + p \<in> float"
    61   unfolding float_def
    62 proof (safe, simp)
    63   fix e1 m1 e2 m2 :: int
    64   { fix e1 m1 e2 m2 :: int assume "e1 \<le> e2"
    65     then have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1"
    66       by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps)
    67     then have "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
    68       by blast }
    69   note * = this
    70   show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
    71   proof (cases e1 e2 rule: linorder_le_cases)
    72     assume "e2 \<le> e1" from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps)
    73   qed (rule *)
    74 qed
    75 
    76 lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float"
    77   apply (auto simp: float_def)
    78   apply hypsubst_thin
    79   apply (rule_tac x="-x" in exI)
    80   apply (rule_tac x="xa" in exI)
    81   apply (simp add: field_simps)
    82   done
    83 
    84 lemma times_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x * y \<in> float"
    85   apply (auto simp: float_def)
    86   apply hypsubst_thin
    87   apply (rule_tac x="x * xa" in exI)
    88   apply (rule_tac x="xb + xc" in exI)
    89   apply (simp add: powr_add)
    90   done
    91 
    92 lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
    93   using plus_float [of x "- y"] by simp
    94 
    95 lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float"
    96   by (cases x rule: linorder_cases[of 0]) auto
    97 
    98 lemma sgn_of_float[simp]: "x \<in> float \<Longrightarrow> sgn x \<in> float"
    99   by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float)
   100 
   101 lemma div_power_2_float[simp]: "x \<in> float \<Longrightarrow> x / 2^d \<in> float"
   102   apply (auto simp add: float_def)
   103   apply hypsubst_thin
   104   apply (rule_tac x="x" in exI)
   105   apply (rule_tac x="xa - d" in exI)
   106   apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])
   107   done
   108 
   109 lemma div_power_2_int_float[simp]: "x \<in> float \<Longrightarrow> x / (2::int)^d \<in> float"
   110   apply (auto simp add: float_def)
   111   apply hypsubst_thin
   112   apply (rule_tac x="x" in exI)
   113   apply (rule_tac x="xa - d" in exI)
   114   apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])
   115   done
   116 
   117 lemma div_numeral_Bit0_float[simp]:
   118   assumes x: "x / numeral n \<in> float" shows "x / (numeral (Num.Bit0 n)) \<in> float"
   119 proof -
   120   have "(x / numeral n) / 2^1 \<in> float"
   121     by (intro x div_power_2_float)
   122   also have "(x / numeral n) / 2^1 = x / (numeral (Num.Bit0 n))"
   123     by (induct n) auto
   124   finally show ?thesis .
   125 qed
   126 
   127 lemma div_neg_numeral_Bit0_float[simp]:
   128   assumes x: "x / numeral n \<in> float" shows "x / (- numeral (Num.Bit0 n)) \<in> float"
   129 proof -
   130   have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp
   131   also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)"
   132     by simp
   133   finally show ?thesis .
   134 qed
   135 
   136 lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e" by simp
   137 declare Float.rep_eq[simp]
   138 
   139 lemma compute_real_of_float[code]:
   140   "real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))"
   141 by (simp add: real_of_float_def[symmetric] powr_int)
   142 
   143 code_datatype Float
   144 
   145 subsection {* Arithmetic operations on floating point numbers *}
   146 
   147 instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}"
   148 begin
   149 
   150 lift_definition zero_float :: float is 0 by simp
   151 declare zero_float.rep_eq[simp]
   152 lift_definition one_float :: float is 1 by simp
   153 declare one_float.rep_eq[simp]
   154 lift_definition plus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op +" by simp
   155 declare plus_float.rep_eq[simp]
   156 lift_definition times_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op *" by simp
   157 declare times_float.rep_eq[simp]
   158 lift_definition minus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op -" by simp
   159 declare minus_float.rep_eq[simp]
   160 lift_definition uminus_float :: "float \<Rightarrow> float" is "uminus" by simp
   161 declare uminus_float.rep_eq[simp]
   162 
   163 lift_definition abs_float :: "float \<Rightarrow> float" is abs by simp
   164 declare abs_float.rep_eq[simp]
   165 lift_definition sgn_float :: "float \<Rightarrow> float" is sgn by simp
   166 declare sgn_float.rep_eq[simp]
   167 
   168 lift_definition equal_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op = :: real \<Rightarrow> real \<Rightarrow> bool" .
   169 
   170 lift_definition less_eq_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op \<le>" .
   171 declare less_eq_float.rep_eq[simp]
   172 lift_definition less_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op <" .
   173 declare less_float.rep_eq[simp]
   174 
   175 instance
   176   proof qed (transfer, fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+
   177 end
   178 
   179 lemma real_of_float_power[simp]: fixes f::float shows "real (f^n) = real f^n"
   180   by (induct n) simp_all
   181 
   182 lemma fixes x y::float
   183   shows real_of_float_min: "real (min x y) = min (real x) (real y)"
   184     and real_of_float_max: "real (max x y) = max (real x) (real y)"
   185   by (simp_all add: min_def max_def)
   186 
   187 instance float :: unbounded_dense_linorder
   188 proof
   189   fix a b :: float
   190   show "\<exists>c. a < c"
   191     apply (intro exI[of _ "a + 1"])
   192     apply transfer
   193     apply simp
   194     done
   195   show "\<exists>c. c < a"
   196     apply (intro exI[of _ "a - 1"])
   197     apply transfer
   198     apply simp
   199     done
   200   assume "a < b"
   201   then show "\<exists>c. a < c \<and> c < b"
   202     apply (intro exI[of _ "(a + b) * Float 1 -1"])
   203     apply transfer
   204     apply (simp add: powr_minus)
   205     done
   206 qed
   207 
   208 instantiation float :: lattice_ab_group_add
   209 begin
   210 
   211 definition inf_float::"float\<Rightarrow>float\<Rightarrow>float"
   212 where "inf_float a b = min a b"
   213 
   214 definition sup_float::"float\<Rightarrow>float\<Rightarrow>float"
   215 where "sup_float a b = max a b"
   216 
   217 instance
   218   by default
   219      (transfer, simp_all add: inf_float_def sup_float_def real_of_float_min real_of_float_max)+
   220 end
   221 
   222 lemma float_numeral[simp]: "real (numeral x :: float) = numeral x"
   223   apply (induct x)
   224   apply simp
   225   apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq real_float
   226                   plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float)
   227   done
   228 
   229 lemma transfer_numeral [transfer_rule]:
   230   "rel_fun (op =) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"
   231   unfolding rel_fun_def float.pcr_cr_eq  cr_float_def by simp
   232 
   233 lemma float_neg_numeral[simp]: "real (- numeral x :: float) = - numeral x"
   234   by simp
   235 
   236 lemma transfer_neg_numeral [transfer_rule]:
   237   "rel_fun (op =) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)"
   238   unfolding rel_fun_def float.pcr_cr_eq cr_float_def by simp
   239 
   240 lemma
   241   shows float_of_numeral[simp]: "numeral k = float_of (numeral k)"
   242     and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)"
   243   unfolding real_of_float_eq by simp_all
   244 
   245 subsection {* Represent floats as unique mantissa and exponent *}
   246 
   247 lemma int_induct_abs[case_names less]:
   248   fixes j :: int
   249   assumes H: "\<And>n. (\<And>i. \<bar>i\<bar> < \<bar>n\<bar> \<Longrightarrow> P i) \<Longrightarrow> P n"
   250   shows "P j"
   251 proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct)
   252   case less show ?case by (rule H[OF less]) simp
   253 qed
   254 
   255 lemma int_cancel_factors:
   256   fixes n :: int assumes "1 < r" shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)"
   257 proof (induct n rule: int_induct_abs)
   258   case (less n)
   259   { fix m assume n: "n \<noteq> 0" "n = m * r"
   260     then have "\<bar>m \<bar> < \<bar>n\<bar>"
   261       by (metis abs_dvd_iff abs_ge_self assms comm_semiring_1_class.normalizing_semiring_rules(7)
   262                 dvd_imp_le_int dvd_refl dvd_triv_right linorder_neq_iff linorder_not_le
   263                 mult_eq_0_iff zdvd_mult_cancel1)
   264     from less[OF this] n have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" by auto }
   265   then show ?case
   266     by (metis comm_semiring_1_class.normalizing_semiring_rules(12,7) dvdE power_0)
   267 qed
   268 
   269 lemma mult_powr_eq_mult_powr_iff_asym:
   270   fixes m1 m2 e1 e2 :: int
   271   assumes m1: "\<not> 2 dvd m1" and "e1 \<le> e2"
   272   shows "m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
   273 proof
   274   have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto
   275   assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2"
   276   with `e1 \<le> e2` have "m1 = m2 * 2 powr nat (e2 - e1)"
   277     by (simp add: powr_divide2[symmetric] field_simps)
   278   also have "\<dots> = m2 * 2^nat (e2 - e1)"
   279     by (simp add: powr_realpow)
   280   finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"
   281     unfolding real_of_int_inject .
   282   with m1 have "m1 = m2"
   283     by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
   284   then show "m1 = m2 \<and> e1 = e2"
   285     using eq `m1 \<noteq> 0` by (simp add: powr_inj)
   286 qed simp
   287 
   288 lemma mult_powr_eq_mult_powr_iff:
   289   fixes m1 m2 e1 e2 :: int
   290   shows "\<not> 2 dvd m1 \<Longrightarrow> \<not> 2 dvd m2 \<Longrightarrow> m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
   291   using mult_powr_eq_mult_powr_iff_asym[of m1 e1 e2 m2]
   292   using mult_powr_eq_mult_powr_iff_asym[of m2 e2 e1 m1]
   293   by (cases e1 e2 rule: linorder_le_cases) auto
   294 
   295 lemma floatE_normed:
   296   assumes x: "x \<in> float"
   297   obtains (zero) "x = 0"
   298    | (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0"
   299 proof atomize_elim
   300   { assume "x \<noteq> 0"
   301     from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def)
   302     with `x \<noteq> 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
   303       by auto
   304     with `\<not> 2 dvd k` x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m"
   305       by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"])
   306          (simp add: powr_add powr_realpow) }
   307   then show "x = 0 \<or> (\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m \<and> x \<noteq> 0)"
   308     by blast
   309 qed
   310 
   311 lemma float_normed_cases:
   312   fixes f :: float
   313   obtains (zero) "f = 0"
   314    | (powr) m e :: int where "real f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0"
   315 proof (atomize_elim, induct f)
   316   case (float_of y) then show ?case
   317     by (cases rule: floatE_normed) (auto simp: zero_float_def)
   318 qed
   319 
   320 definition mantissa :: "float \<Rightarrow> int" where
   321   "mantissa f = fst (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
   322    \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))"
   323 
   324 definition exponent :: "float \<Rightarrow> int" where
   325   "exponent f = snd (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
   326    \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))"
   327 
   328 lemma
   329   shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E)
   330     and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M)
   331 proof -
   332   have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)" by auto
   333   then show ?E ?M
   334     by (auto simp add: mantissa_def exponent_def zero_float_def)
   335 qed
   336 
   337 lemma
   338   shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E)
   339     and mantissa_not_dvd: "f \<noteq> (float_of 0) \<Longrightarrow> \<not> 2 dvd mantissa f" (is "_ \<Longrightarrow> ?D")
   340 proof cases
   341   assume [simp]: "f \<noteq> (float_of 0)"
   342   have "f = mantissa f * 2 powr exponent f \<and> \<not> 2 dvd mantissa f"
   343   proof (cases f rule: float_normed_cases)
   344     case (powr m e)
   345     then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
   346      \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p)"
   347       by auto
   348     then show ?thesis
   349       unfolding exponent_def mantissa_def
   350       by (rule someI2_ex) (simp add: zero_float_def)
   351   qed (simp add: zero_float_def)
   352   then show ?E ?D by auto
   353 qed simp
   354 
   355 lemma mantissa_noteq_0: "f \<noteq> float_of 0 \<Longrightarrow> mantissa f \<noteq> 0"
   356   using mantissa_not_dvd[of f] by auto
   357 
   358 lemma
   359   fixes m e :: int
   360   defines "f \<equiv> float_of (m * 2 powr e)"
   361   assumes dvd: "\<not> 2 dvd m"
   362   shows mantissa_float: "mantissa f = m" (is "?M")
   363     and exponent_float: "m \<noteq> 0 \<Longrightarrow> exponent f = e" (is "_ \<Longrightarrow> ?E")
   364 proof cases
   365   assume "m = 0" with dvd show "mantissa f = m" by auto
   366 next
   367   assume "m \<noteq> 0"
   368   then have f_not_0: "f \<noteq> float_of 0" by (simp add: f_def)
   369   from mantissa_exponent[of f]
   370   have "m * 2 powr e = mantissa f * 2 powr exponent f"
   371     by (auto simp add: f_def)
   372   then show "?M" "?E"
   373     using mantissa_not_dvd[OF f_not_0] dvd
   374     by (auto simp: mult_powr_eq_mult_powr_iff)
   375 qed
   376 
   377 subsection {* Compute arithmetic operations *}
   378 
   379 lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"
   380   unfolding real_of_float_eq mantissa_exponent[of f] by simp
   381 
   382 lemma Float_cases[case_names Float, cases type: float]:
   383   fixes f :: float
   384   obtains (Float) m e :: int where "f = Float m e"
   385   using Float_mantissa_exponent[symmetric]
   386   by (atomize_elim) auto
   387 
   388 lemma denormalize_shift:
   389   assumes f_def: "f \<equiv> Float m e" and not_0: "f \<noteq> float_of 0"
   390   obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i"
   391 proof
   392   from mantissa_exponent[of f] f_def
   393   have "m * 2 powr e = mantissa f * 2 powr exponent f"
   394     by simp
   395   then have eq: "m = mantissa f * 2 powr (exponent f - e)"
   396     by (simp add: powr_divide2[symmetric] field_simps)
   397   moreover
   398   have "e \<le> exponent f"
   399   proof (rule ccontr)
   400     assume "\<not> e \<le> exponent f"
   401     then have pos: "exponent f < e" by simp
   402     then have "2 powr (exponent f - e) = 2 powr - real (e - exponent f)"
   403       by simp
   404     also have "\<dots> = 1 / 2^nat (e - exponent f)"
   405       using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric])
   406     finally have "m * 2^nat (e - exponent f) = real (mantissa f)"
   407       using eq by simp
   408     then have "mantissa f = m * 2^nat (e - exponent f)"
   409       unfolding real_of_int_inject by simp
   410     with `exponent f < e` have "2 dvd mantissa f"
   411       apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"])
   412       apply (cases "nat (e - exponent f)")
   413       apply auto
   414       done
   415     then show False using mantissa_not_dvd[OF not_0] by simp
   416   qed
   417   ultimately have "real m = mantissa f * 2^nat (exponent f - e)"
   418     by (simp add: powr_realpow[symmetric])
   419   with `e \<le> exponent f`
   420   show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)"
   421     unfolding real_of_int_inject by auto
   422 qed
   423 
   424 lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0"
   425   by transfer simp
   426 hide_fact (open) compute_float_zero
   427 
   428 lemma compute_float_one[code_unfold, code]: "1 = Float 1 0"
   429   by transfer simp
   430 hide_fact (open) compute_float_one
   431 
   432 definition normfloat :: "float \<Rightarrow> float" where
   433   [simp]: "normfloat x = x"
   434 
   435 lemma compute_normfloat[code]: "normfloat (Float m e) =
   436   (if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1))
   437                            else if m = 0 then 0 else Float m e)"
   438   unfolding normfloat_def
   439   by transfer (auto simp add: powr_add zmod_eq_0_iff)
   440 hide_fact (open) compute_normfloat
   441 
   442 lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k"
   443   by transfer simp
   444 hide_fact (open) compute_float_numeral
   445 
   446 lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k"
   447   by transfer simp
   448 hide_fact (open) compute_float_neg_numeral
   449 
   450 lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1"
   451   by transfer simp
   452 hide_fact (open) compute_float_uminus
   453 
   454 lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)"
   455   by transfer (simp add: field_simps powr_add)
   456 hide_fact (open) compute_float_times
   457 
   458 lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 =
   459   (if m1 = 0 then Float m2 e2 else if m2 = 0 then Float m1 e1 else
   460   if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1
   461               else Float (m2 + m1 * 2^nat (e1 - e2)) e2)"
   462   by transfer (simp add: field_simps powr_realpow[symmetric] powr_divide2[symmetric])
   463 hide_fact (open) compute_float_plus
   464 
   465 lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)"
   466   by simp
   467 hide_fact (open) compute_float_minus
   468 
   469 lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"
   470   by transfer (simp add: sgn_times)
   471 hide_fact (open) compute_float_sgn
   472 
   473 lift_definition is_float_pos :: "float \<Rightarrow> bool" is "op < 0 :: real \<Rightarrow> bool" .
   474 
   475 lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m"
   476   by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0])
   477 hide_fact (open) compute_is_float_pos
   478 
   479 lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)"
   480   by transfer (simp add: field_simps)
   481 hide_fact (open) compute_float_less
   482 
   483 lift_definition is_float_nonneg :: "float \<Rightarrow> bool" is "op \<le> 0 :: real \<Rightarrow> bool" .
   484 
   485 lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m"
   486   by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0])
   487 hide_fact (open) compute_is_float_nonneg
   488 
   489 lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)"
   490   by transfer (simp add: field_simps)
   491 hide_fact (open) compute_float_le
   492 
   493 lift_definition is_float_zero :: "float \<Rightarrow> bool"  is "op = 0 :: real \<Rightarrow> bool" .
   494 
   495 lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m"
   496   by transfer (auto simp add: is_float_zero_def)
   497 hide_fact (open) compute_is_float_zero
   498 
   499 lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e"
   500   by transfer (simp add: abs_mult)
   501 hide_fact (open) compute_float_abs
   502 
   503 lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)"
   504   by transfer simp
   505 hide_fact (open) compute_float_eq
   506 
   507 subsection {* Rounding Real numbers *}
   508 
   509 definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where
   510   "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec"
   511 
   512 definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" where
   513   "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec"
   514 
   515 lemma round_down_float[simp]: "round_down prec x \<in> float"
   516   unfolding round_down_def
   517   by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus)
   518 
   519 lemma round_up_float[simp]: "round_up prec x \<in> float"
   520   unfolding round_up_def
   521   by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus)
   522 
   523 lemma round_up: "x \<le> round_up prec x"
   524   by (simp add: powr_minus_divide le_divide_eq round_up_def)
   525 
   526 lemma round_down: "round_down prec x \<le> x"
   527   by (simp add: powr_minus_divide divide_le_eq round_down_def)
   528 
   529 lemma round_up_0[simp]: "round_up p 0 = 0"
   530   unfolding round_up_def by simp
   531 
   532 lemma round_down_0[simp]: "round_down p 0 = 0"
   533   unfolding round_down_def by simp
   534 
   535 lemma round_up_diff_round_down:
   536   "round_up prec x - round_down prec x \<le> 2 powr -prec"
   537 proof -
   538   have "round_up prec x - round_down prec x =
   539     (ceiling (x * 2 powr prec) - floor (x * 2 powr prec)) * 2 powr -prec"
   540     by (simp add: round_up_def round_down_def field_simps)
   541   also have "\<dots> \<le> 1 * 2 powr -prec"
   542     by (rule mult_mono)
   543        (auto simp del: real_of_int_diff
   544              simp: real_of_int_diff[symmetric] real_of_int_le_one_cancel_iff ceiling_diff_floor_le_1)
   545   finally show ?thesis by simp
   546 qed
   547 
   548 lemma round_down_shift: "round_down p (x * 2 powr k) = 2 powr k * round_down (p + k) x"
   549   unfolding round_down_def
   550   by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric])
   551     (simp add: powr_add[symmetric])
   552 
   553 lemma round_up_shift: "round_up p (x * 2 powr k) = 2 powr k * round_up (p + k) x"
   554   unfolding round_up_def
   555   by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric])
   556     (simp add: powr_add[symmetric])
   557 
   558 subsection {* Rounding Floats *}
   559 
   560 lift_definition float_up :: "int \<Rightarrow> float \<Rightarrow> float" is round_up by simp
   561 declare float_up.rep_eq[simp]
   562 
   563 lemma round_up_correct:
   564   shows "round_up e f - f \<in> {0..2 powr -e}"
   565 unfolding atLeastAtMost_iff
   566 proof
   567   have "round_up e f - f \<le> round_up e f - round_down e f" using round_down by simp
   568   also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
   569   finally show "round_up e f - f \<le> 2 powr real (- e)"
   570     by simp
   571 qed (simp add: algebra_simps round_up)
   572 
   573 lemma float_up_correct:
   574   shows "real (float_up e f) - real f \<in> {0..2 powr -e}"
   575   by transfer (rule round_up_correct)
   576 
   577 lift_definition float_down :: "int \<Rightarrow> float \<Rightarrow> float" is round_down by simp
   578 declare float_down.rep_eq[simp]
   579 
   580 lemma round_down_correct:
   581   shows "f - (round_down e f) \<in> {0..2 powr -e}"
   582 unfolding atLeastAtMost_iff
   583 proof
   584   have "f - round_down e f \<le> round_up e f - round_down e f" using round_up by simp
   585   also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
   586   finally show "f - round_down e f \<le> 2 powr real (- e)"
   587     by simp
   588 qed (simp add: algebra_simps round_down)
   589 
   590 lemma float_down_correct:
   591   shows "real f - real (float_down e f) \<in> {0..2 powr -e}"
   592   by transfer (rule round_down_correct)
   593 
   594 lemma compute_float_down[code]:
   595   "float_down p (Float m e) =
   596     (if p + e < 0 then Float (m div 2^nat (-(p + e))) (-p) else Float m e)"
   597 proof cases
   598   assume "p + e < 0"
   599   hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
   600     using powr_realpow[of 2 "nat (-(p + e))"] by simp
   601   also have "... = 1 / 2 powr p / 2 powr e"
   602     unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
   603   finally show ?thesis
   604     using `p + e < 0`
   605     by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric])
   606 next
   607   assume "\<not> p + e < 0"
   608   then have r: "real e + real p = real (nat (e + p))" by simp
   609   have r: "\<lfloor>(m * 2 powr e) * 2 powr real p\<rfloor> = (m * 2 powr e) * 2 powr real p"
   610     by (auto intro: exI[where x="m*2^nat (e+p)"]
   611              simp add: ac_simps powr_add[symmetric] r powr_realpow)
   612   with `\<not> p + e < 0` show ?thesis
   613     by transfer
   614        (auto simp add: round_down_def field_simps powr_add powr_minus inverse_eq_divide)
   615 qed
   616 hide_fact (open) compute_float_down
   617 
   618 lemma abs_round_down_le: "\<bar>f - (round_down e f)\<bar> \<le> 2 powr -e"
   619   using round_down_correct[of f e] by simp
   620 
   621 lemma abs_round_up_le: "\<bar>f - (round_up e f)\<bar> \<le> 2 powr -e"
   622   using round_up_correct[of e f] by simp
   623 
   624 lemma round_down_nonneg: "0 \<le> s \<Longrightarrow> 0 \<le> round_down p s"
   625   by (auto simp: round_down_def)
   626 
   627 lemma ceil_divide_floor_conv:
   628 assumes "b \<noteq> 0"
   629 shows "\<lceil>real a / real b\<rceil> = (if b dvd a then a div b else \<lfloor>real a / real b\<rfloor> + 1)"
   630 proof cases
   631   assume "\<not> b dvd a"
   632   hence "a mod b \<noteq> 0" by auto
   633   hence ne: "real (a mod b) / real b \<noteq> 0" using `b \<noteq> 0` by auto
   634   have "\<lceil>real a / real b\<rceil> = \<lfloor>real a / real b\<rfloor> + 1"
   635   apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric])
   636   proof -
   637     have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b" by simp
   638     moreover have "real \<lfloor>real a / real b\<rfloor> \<noteq> real a / real b"
   639     apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \<noteq> 0` by auto
   640     ultimately show "real \<lfloor>real a / real b\<rfloor> < real a / real b" by arith
   641   qed
   642   thus ?thesis using `\<not> b dvd a` by simp
   643 qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]
   644   floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)
   645 
   646 lemma compute_float_up[code]:
   647   "float_up p (Float m e) =
   648     (let P = 2^nat (-(p + e)); r = m mod P in
   649       if p + e < 0 then Float (m div P + (if r = 0 then 0 else 1)) (-p) else Float m e)"
   650 proof cases
   651   assume "p + e < 0"
   652   hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
   653     using powr_realpow[of 2 "nat (-(p + e))"] by simp
   654   also have "... = 1 / 2 powr p / 2 powr e"
   655   unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
   656   finally have twopow_rewrite:
   657     "real ((2::int) ^ nat (- (p + e))) = 1 / 2 powr real p / 2 powr real e" .
   658   with `p + e < 0` have powr_rewrite:
   659     "2 powr real e * 2 powr real p = 1 / real ((2::int) ^ nat (- (p + e)))"
   660     unfolding powr_divide2 by simp
   661   show ?thesis
   662   proof cases
   663     assume "2^nat (-(p + e)) dvd m"
   664     with `p + e < 0` twopow_rewrite show ?thesis
   665       by transfer (auto simp: ac_simps round_up_def floor_divide_eq_div dvd_eq_mod_eq_0)
   666   next
   667     assume ndvd: "\<not> 2 ^ nat (- (p + e)) dvd m"
   668     have one_div: "real m * (1 / real ((2::int) ^ nat (- (p + e)))) =
   669       real m / real ((2::int) ^ nat (- (p + e)))"
   670       by (simp add: field_simps)
   671     have "real \<lceil>real m * (2 powr real e * 2 powr real p)\<rceil> =
   672       real \<lfloor>real m * (2 powr real e * 2 powr real p)\<rfloor> + 1"
   673       using ndvd unfolding powr_rewrite one_div
   674       by (subst ceil_divide_floor_conv) (auto simp: field_simps)
   675     thus ?thesis using `p + e < 0` twopow_rewrite
   676       by transfer (auto simp: ac_simps round_up_def floor_divide_eq_div[symmetric])
   677   qed
   678 next
   679   assume "\<not> p + e < 0"
   680   then have r1: "real e + real p = real (nat (e + p))" by simp
   681   have r: "\<lceil>(m * 2 powr e) * 2 powr real p\<rceil> = (m * 2 powr e) * 2 powr real p"
   682     by (auto simp add: ac_simps powr_add[symmetric] r1 powr_realpow
   683       intro: exI[where x="m*2^nat (e+p)"])
   684   then show ?thesis using `\<not> p + e < 0`
   685     by transfer
   686        (simp add: round_up_def floor_divide_eq_div field_simps powr_add powr_minus inverse_eq_divide)
   687 qed
   688 hide_fact (open) compute_float_up
   689 
   690 lemmas real_of_ints =
   691   real_of_int_zero
   692   real_of_one
   693   real_of_int_add
   694   real_of_int_minus
   695   real_of_int_diff
   696   real_of_int_mult
   697   real_of_int_power
   698   real_numeral
   699 lemmas real_of_nats =
   700   real_of_nat_zero
   701   real_of_nat_one
   702   real_of_nat_1
   703   real_of_nat_add
   704   real_of_nat_mult
   705   real_of_nat_power
   706 
   707 lemmas int_of_reals = real_of_ints[symmetric]
   708 lemmas nat_of_reals = real_of_nats[symmetric]
   709 
   710 lemma two_real_int: "(2::real) = real (2::int)" by simp
   711 lemma two_real_nat: "(2::real) = real (2::nat)" by simp
   712 
   713 lemma mult_cong: "a = c ==> b = d ==> a*b = c*d" by simp
   714 
   715 subsection {* Compute bitlen of integers *}
   716 
   717 definition bitlen :: "int \<Rightarrow> int" where
   718   "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)"
   719 
   720 lemma bitlen_nonneg: "0 \<le> bitlen x"
   721 proof -
   722   {
   723     assume "0 > x"
   724     have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all
   725     also have "... < log 2 (-x)" using `0 > x` by auto
   726     finally have "-1 < log 2 (-x)" .
   727   } thus "0 \<le> bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg)
   728 qed
   729 
   730 lemma bitlen_bounds:
   731   assumes "x > 0"
   732   shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)"
   733 proof
   734   have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))"
   735     using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] `x > 0`
   736     using real_nat_eq_real[of "floor (log 2 (real x))"]
   737     by simp
   738   also have "... \<le> 2 powr log 2 (real x)"
   739     by simp
   740   also have "... = real x"
   741     using `0 < x` by simp
   742   finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x" by simp
   743   thus "2 ^ nat (bitlen x - 1) \<le> x" using `x > 0`
   744     by (simp add: bitlen_def)
   745 next
   746   have "x \<le> 2 powr (log 2 x)" using `x > 0` by simp
   747   also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"
   748     apply (simp add: powr_realpow[symmetric])
   749     using `x > 0` by simp
   750   finally show "x < 2 ^ nat (bitlen x)" using `x > 0`
   751     by (simp add: bitlen_def ac_simps int_of_reals del: real_of_ints)
   752 qed
   753 
   754 lemma bitlen_pow2[simp]:
   755   assumes "b > 0"
   756   shows "bitlen (b * 2 ^ c) = bitlen b + c"
   757 proof -
   758   from assms have "b * 2 ^ c > 0" by auto
   759   thus ?thesis
   760     using floor_add[of "log 2 b" c] assms
   761     by (auto simp add: log_mult log_nat_power bitlen_def)
   762 qed
   763 
   764 lemma bitlen_Float:
   765   fixes m e
   766   defines "f \<equiv> Float m e"
   767   shows "bitlen (\<bar>mantissa f\<bar>) + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)"
   768 proof (cases "m = 0")
   769   case True
   770   then show ?thesis by (simp add: f_def bitlen_def Float_def)
   771 next
   772   case False
   773   hence "f \<noteq> float_of 0"
   774     unfolding real_of_float_eq by (simp add: f_def)
   775   hence "mantissa f \<noteq> 0"
   776     by (simp add: mantissa_noteq_0)
   777   moreover
   778   obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i"
   779     by (rule f_def[THEN denormalize_shift, OF `f \<noteq> float_of 0`])
   780   ultimately show ?thesis by (simp add: abs_mult)
   781 qed
   782 
   783 lemma compute_bitlen[code]:
   784   shows "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)"
   785 proof -
   786   { assume "2 \<le> x"
   787     then have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 (x - x mod 2)\<rfloor>"
   788       by (simp add: log_mult zmod_zdiv_equality')
   789     also have "\<dots> = \<lfloor>log 2 (real x)\<rfloor>"
   790     proof cases
   791       assume "x mod 2 = 0" then show ?thesis by simp
   792     next
   793       def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>"
   794       then have "0 \<le> n"
   795         using `2 \<le> x` by simp
   796       assume "x mod 2 \<noteq> 0"
   797       with `2 \<le> x` have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0)
   798       with `2 \<le> x` have "x \<noteq> 2^nat n" by (cases "nat n") auto
   799       moreover
   800       { have "real (2^nat n :: int) = 2 powr (nat n)"
   801           by (simp add: powr_realpow)
   802         also have "\<dots> \<le> 2 powr (log 2 x)"
   803           using `2 \<le> x` by (simp add: n_def del: powr_log_cancel)
   804         finally have "2^nat n \<le> x" using `2 \<le> x` by simp }
   805       ultimately have "2^nat n \<le> x - 1" by simp
   806       then have "2^nat n \<le> real (x - 1)"
   807         unfolding real_of_int_le_iff[symmetric] by simp
   808       { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
   809           using `0 \<le> n` by (simp add: log_nat_power)
   810         also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>"
   811           using `2^nat n \<le> real (x - 1)` `0 \<le> n` `2 \<le> x` by (auto intro: floor_mono)
   812         finally have "n \<le> \<lfloor>log 2 (x - 1)\<rfloor>" . }
   813       moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n"
   814         using `2 \<le> x` by (auto simp add: n_def intro!: floor_mono)
   815       ultimately show "\<lfloor>log 2 (x - x mod 2)\<rfloor> = \<lfloor>log 2 x\<rfloor>"
   816         unfolding n_def `x mod 2 = 1` by auto
   817     qed
   818     finally have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 x\<rfloor>" . }
   819   moreover
   820   { assume "x < 2" "0 < x"
   821     then have "x = 1" by simp
   822     then have "\<lfloor>log 2 (real x)\<rfloor> = 0" by simp }
   823   ultimately show ?thesis
   824     unfolding bitlen_def
   825     by (auto simp: pos_imp_zdiv_pos_iff not_le)
   826 qed
   827 hide_fact (open) compute_bitlen
   828 
   829 lemma float_gt1_scale: assumes "1 \<le> Float m e"
   830   shows "0 \<le> e + (bitlen m - 1)"
   831 proof -
   832   have "0 < Float m e" using assms by auto
   833   hence "0 < m" using powr_gt_zero[of 2 e]
   834     by (auto simp: zero_less_mult_iff)
   835   hence "m \<noteq> 0" by auto
   836   show ?thesis
   837   proof (cases "0 \<le> e")
   838     case True thus ?thesis using `0 < m`  by (simp add: bitlen_def)
   839   next
   840     have "(1::int) < 2" by simp
   841     case False let ?S = "2^(nat (-e))"
   842     have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"]
   843       by (auto simp: powr_minus field_simps inverse_eq_divide)
   844     hence "1 \<le> real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"]
   845       by (auto simp: powr_minus)
   846     hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)
   847     hence "?S \<le> real m" unfolding mult_assoc by auto
   848     hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto
   849     from this bitlen_bounds[OF `0 < m`, THEN conjunct2]
   850     have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric] by (rule order_le_less_trans)
   851     hence "-e < bitlen m" using False by auto
   852     thus ?thesis by auto
   853   qed
   854 qed
   855 
   856 lemma bitlen_div: assumes "0 < m" shows "1 \<le> real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2"
   857 proof -
   858   let ?B = "2^nat(bitlen m - 1)"
   859 
   860   have "?B \<le> m" using bitlen_bounds[OF `0 <m`] ..
   861   hence "1 * ?B \<le> real m" unfolding real_of_int_le_iff[symmetric] by auto
   862   thus "1 \<le> real m / ?B" by auto
   863 
   864   have "m \<noteq> 0" using assms by auto
   865   have "0 \<le> bitlen m - 1" using `0 < m` by (auto simp: bitlen_def)
   866 
   867   have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 <m`] ..
   868   also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using `0 < m` by (auto simp: bitlen_def)
   869   also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF `0 \<le> bitlen m - 1` zero_le_one] by auto
   870   finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto
   871   hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto)
   872   thus "real m / ?B < 2" by auto
   873 qed
   874 
   875 subsection {* Approximation of positive rationals *}
   876 
   877 lemma zdiv_zmult_twopow_eq: fixes a b::int shows "a div b div (2 ^ n) = a div (b * 2 ^ n)"
   878 by (simp add: zdiv_zmult2_eq)
   879 
   880 lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)"
   881   by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps)
   882 
   883 lemma real_div_nat_eq_floor_of_divide:
   884   fixes a b::nat
   885   shows "a div b = real (floor (a/b))"
   886 by (metis floor_divide_eq_div real_of_int_of_nat_eq zdiv_int)
   887 
   888 definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)"
   889 
   890 lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
   891   is "\<lambda>prec (x::nat) (y::nat). round_down (rat_precision prec x y) (x / y)" by simp
   892 
   893 lemma compute_lapprox_posrat[code]:
   894   fixes prec x y
   895   shows "lapprox_posrat prec x y =
   896    (let
   897        l = rat_precision prec x y;
   898        d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y
   899     in normfloat (Float d (- l)))"
   900     unfolding div_mult_twopow_eq normfloat_def
   901     by transfer
   902        (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def
   903              del: two_powr_minus_int_float)
   904 hide_fact (open) compute_lapprox_posrat
   905 
   906 lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
   907   is "\<lambda>prec (x::nat) (y::nat). round_up (rat_precision prec x y) (x / y)" by simp
   908 
   909 (* TODO: optimize using zmod_zmult2_eq, pdivmod ? *)
   910 lemma compute_rapprox_posrat[code]:
   911   fixes prec x y
   912   defines "l \<equiv> rat_precision prec x y"
   913   shows "rapprox_posrat prec x y = (let
   914      l = l ;
   915      X = if 0 \<le> l then (x * 2^nat l, y) else (x, y * 2^nat(-l)) ;
   916      d = fst X div snd X ;
   917      m = fst X mod snd X
   918    in normfloat (Float (d + (if m = 0 \<or> y = 0 then 0 else 1)) (- l)))"
   919 proof (cases "y = 0")
   920   assume "y = 0" thus ?thesis unfolding normfloat_def by transfer simp
   921 next
   922   assume "y \<noteq> 0"
   923   show ?thesis
   924   proof (cases "0 \<le> l")
   925     assume "0 \<le> l"
   926     def x' \<equiv> "x * 2 ^ nat l"
   927     have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power)
   928     moreover have "real x * 2 powr real l = real x'"
   929       by (simp add: powr_realpow[symmetric] `0 \<le> l` x'_def)
   930     ultimately show ?thesis
   931       unfolding normfloat_def
   932       using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \<le> l` `y \<noteq> 0`
   933         l_def[symmetric, THEN meta_eq_to_obj_eq]
   934       by transfer
   935          (simp add: floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0 round_up_def)
   936    next
   937     assume "\<not> 0 \<le> l"
   938     def y' \<equiv> "y * 2 ^ nat (- l)"
   939     from `y \<noteq> 0` have "y' \<noteq> 0" by (simp add: y'_def)
   940     have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power)
   941     moreover have "real x * real (2::int) powr real l / real y = x / real y'"
   942       using `\<not> 0 \<le> l`
   943       by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps inverse_eq_divide)
   944     ultimately show ?thesis
   945       unfolding normfloat_def
   946       using ceil_divide_floor_conv[of y' x] `\<not> 0 \<le> l` `y' \<noteq> 0` `y \<noteq> 0`
   947         l_def[symmetric, THEN meta_eq_to_obj_eq]
   948       by transfer
   949          (simp add: round_up_def ceil_divide_floor_conv floor_divide_eq_div[symmetric] dvd_eq_mod_eq_0)
   950   qed
   951 qed
   952 hide_fact (open) compute_rapprox_posrat
   953 
   954 lemma rat_precision_pos:
   955   assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n"
   956   shows "rat_precision n (int x) (int y) > 0"
   957 proof -
   958   { assume "0 < x" hence "log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) }
   959   hence "bitlen (int x) < bitlen (int y)" using assms
   960     by (simp add: bitlen_def del: floor_add_one)
   961       (auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one)
   962   thus ?thesis
   963     using assms by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def)
   964 qed
   965 
   966 lemma power_aux:
   967   assumes "x > 0"
   968   shows "(2::int) ^ nat (x - 1) \<le> 2 ^ nat x - 1"
   969 proof -
   970   def y \<equiv> "nat (x - 1)"
   971   moreover
   972   have "(2::int) ^ y \<le> (2 ^ (y + 1)) - 1" by simp
   973   ultimately show ?thesis using assms by simp
   974 qed
   975 
   976 lemma rapprox_posrat_less1:
   977   assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n"
   978   shows "real (rapprox_posrat n x y) < 1"
   979 proof -
   980   have powr1: "2 powr real (rat_precision n (int x) (int y)) =
   981     2 ^ nat (rat_precision n (int x) (int y))" using rat_precision_pos[of x y n] assms
   982     by (simp add: powr_realpow[symmetric])
   983   have "x * 2 powr real (rat_precision n (int x) (int y)) / y = (x / y) *
   984      2 powr real (rat_precision n (int x) (int y))" by simp
   985   also have "... < (1 / 2) * 2 powr real (rat_precision n (int x) (int y))"
   986     apply (rule mult_strict_right_mono) by (insert assms) auto
   987   also have "\<dots> = 2 powr real (rat_precision n (int x) (int y) - 1)"
   988     using powr_add [of 2 _ "- 1", simplified add_uminus_conv_diff] by (simp add: powr_minus)
   989   also have "\<dots> = 2 ^ nat (rat_precision n (int x) (int y) - 1)"
   990     using rat_precision_pos[of x y n] assms by (simp add: powr_realpow[symmetric])
   991   also have "\<dots> \<le> 2 ^ nat (rat_precision n (int x) (int y)) - 1"
   992     unfolding int_of_reals real_of_int_le_iff
   993     using rat_precision_pos[OF assms] by (rule power_aux)
   994   finally show ?thesis
   995     apply (transfer fixing: n x y)
   996     apply (simp add: round_up_def field_simps powr_minus inverse_eq_divide powr1)
   997     unfolding int_of_reals real_of_int_less_iff
   998     apply (simp add: ceiling_less_eq)
   999     done
  1000 qed
  1001 
  1002 lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
  1003   "\<lambda>prec (x::int) (y::int). round_down (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp
  1004 
  1005 lemma compute_lapprox_rat[code]:
  1006   "lapprox_rat prec x y =
  1007     (if y = 0 then 0
  1008     else if 0 \<le> x then
  1009       (if 0 < y then lapprox_posrat prec (nat x) (nat y)
  1010       else - (rapprox_posrat prec (nat x) (nat (-y))))
  1011       else (if 0 < y
  1012         then - (rapprox_posrat prec (nat (-x)) (nat y))
  1013         else lapprox_posrat prec (nat (-x)) (nat (-y))))"
  1014   by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps)
  1015 hide_fact (open) compute_lapprox_rat
  1016 
  1017 lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
  1018   "\<lambda>prec (x::int) (y::int). round_up (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp
  1019 
  1020 lemma compute_rapprox_rat[code]:
  1021   "rapprox_rat prec x y =
  1022     (if y = 0 then 0
  1023     else if 0 \<le> x then
  1024       (if 0 < y then rapprox_posrat prec (nat x) (nat y)
  1025       else - (lapprox_posrat prec (nat x) (nat (-y))))
  1026       else (if 0 < y
  1027         then - (lapprox_posrat prec (nat (-x)) (nat y))
  1028         else rapprox_posrat prec (nat (-x)) (nat (-y))))"
  1029   by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps)
  1030 hide_fact (open) compute_rapprox_rat
  1031 
  1032 subsection {* Division *}
  1033 
  1034 definition "real_divl prec a b = round_down (int prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)"
  1035 
  1036 definition "real_divr prec a b = round_up (int prec + \<lfloor> log 2 \<bar>b\<bar> \<rfloor> - \<lfloor> log 2 \<bar>a\<bar> \<rfloor>) (a / b)"
  1037 
  1038 lift_definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divl
  1039   by (simp add: real_divl_def)
  1040 
  1041 lemma compute_float_divl[code]:
  1042   "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
  1043 proof cases
  1044   let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2"
  1045   let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)"
  1046   assume not_0: "m1 \<noteq> 0 \<and> m2 \<noteq> 0"
  1047   then have eq2: "(int prec + \<lfloor>log 2 \<bar>?f2\<bar>\<rfloor> - \<lfloor>log 2 \<bar>?f1\<bar>\<rfloor>) = rat_precision prec \<bar>m1\<bar> \<bar>m2\<bar> + (s2 - s1)"
  1048     by (simp add: abs_mult log_mult rat_precision_def bitlen_def)
  1049   have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s"
  1050     by (simp add: field_simps powr_divide2[symmetric])
  1051 
  1052   show ?thesis
  1053     using not_0
  1054     by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift real_divl_def,
  1055       simp add: field_simps)
  1056 qed (transfer, auto simp: real_divl_def)
  1057 hide_fact (open) compute_float_divl
  1058 
  1059 lift_definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divr
  1060   by (simp add: real_divr_def)
  1061 
  1062 lemma compute_float_divr[code]:
  1063   "float_divr prec (Float m1 s1) (Float m2 s2) = rapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
  1064 proof cases
  1065   let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2"
  1066   let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)"
  1067   assume not_0: "m1 \<noteq> 0 \<and> m2 \<noteq> 0"
  1068   then have eq2: "(int prec + \<lfloor>log 2 \<bar>?f2\<bar>\<rfloor> - \<lfloor>log 2 \<bar>?f1\<bar>\<rfloor>) = rat_precision prec \<bar>m1\<bar> \<bar>m2\<bar> + (s2 - s1)"
  1069     by (simp add: abs_mult log_mult rat_precision_def bitlen_def)
  1070   have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s"
  1071     by (simp add: field_simps powr_divide2[symmetric])
  1072 
  1073   show ?thesis
  1074     using not_0
  1075     by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_up_shift real_divr_def,
  1076       simp add: field_simps)
  1077 qed (transfer, auto simp: real_divr_def)
  1078 hide_fact (open) compute_float_divr
  1079 
  1080 subsection {* Lemmas needed by Approximate *}
  1081 
  1082 lemma Float_num[simp]: shows
  1083    "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and
  1084    "real (Float 1 -1) = 1/2" and "real (Float 1 -2) = 1/4" and "real (Float 1 -3) = 1/8" and
  1085    "real (Float -1 0) = -1" and "real (Float (number_of n) 0) = number_of n"
  1086 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"]
  1087 using powr_realpow[of 2 2] powr_realpow[of 2 3]
  1088 using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3]
  1089 by auto
  1090 
  1091 lemma real_of_Float_int[simp]: "real (Float n 0) = real n" by simp
  1092 
  1093 lemma float_zero[simp]: "real (Float 0 e) = 0" by simp
  1094 
  1095 lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
  1096 by arith
  1097 
  1098 lemma lapprox_rat:
  1099   shows "real (lapprox_rat prec x y) \<le> real x / real y"
  1100   using round_down by (simp add: lapprox_rat_def)
  1101 
  1102 lemma mult_div_le: fixes a b:: int assumes "b > 0" shows "a \<ge> b * (a div b)"
  1103 proof -
  1104   from zmod_zdiv_equality'[of a b]
  1105   have "a = b * (a div b) + a mod b" by simp
  1106   also have "... \<ge> b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign)
  1107   using assms by simp
  1108   finally show ?thesis by simp
  1109 qed
  1110 
  1111 lemma lapprox_rat_nonneg:
  1112   fixes n x y
  1113   defines "p \<equiv> int n - ((bitlen \<bar>x\<bar>) - (bitlen \<bar>y\<bar>))"
  1114   assumes "0 \<le> x" and "0 < y"
  1115   shows "0 \<le> real (lapprox_rat n x y)"
  1116 using assms unfolding lapprox_rat_def p_def[symmetric] round_down_def real_of_int_minus[symmetric]
  1117    powr_int[of 2, simplified]
  1118   by auto
  1119 
  1120 lemma rapprox_rat: "real x / real y \<le> real (rapprox_rat prec x y)"
  1121   using round_up by (simp add: rapprox_rat_def)
  1122 
  1123 lemma rapprox_rat_le1:
  1124   fixes n x y
  1125   assumes xy: "0 \<le> x" "0 < y" "x \<le> y"
  1126   shows "real (rapprox_rat n x y) \<le> 1"
  1127 proof -
  1128   have "bitlen \<bar>x\<bar> \<le> bitlen \<bar>y\<bar>"
  1129     using xy unfolding bitlen_def by (auto intro!: floor_mono)
  1130   then have "0 \<le> rat_precision n \<bar>x\<bar> \<bar>y\<bar>" by (simp add: rat_precision_def)
  1131   have "real \<lceil>real x / real y * 2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)\<rceil>
  1132       \<le> real \<lceil>2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)\<rceil>"
  1133     using xy by (auto intro!: ceiling_mono simp: field_simps)
  1134   also have "\<dots> = 2 powr real (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)"
  1135     using `0 \<le> rat_precision n \<bar>x\<bar> \<bar>y\<bar>`
  1136     by (auto intro!: exI[of _ "2^nat (rat_precision n \<bar>x\<bar> \<bar>y\<bar>)"] simp: powr_int)
  1137   finally show ?thesis
  1138     by (simp add: rapprox_rat_def round_up_def)
  1139        (simp add: powr_minus inverse_eq_divide)
  1140 qed
  1141 
  1142 lemma rapprox_rat_nonneg_neg:
  1143   "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
  1144   unfolding rapprox_rat_def round_up_def
  1145   by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff)
  1146 
  1147 lemma rapprox_rat_neg:
  1148   "x < 0 \<Longrightarrow> 0 < y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
  1149   unfolding rapprox_rat_def round_up_def
  1150   by (auto simp: field_simps mult_le_0_iff)
  1151 
  1152 lemma rapprox_rat_nonpos_pos:
  1153   "x \<le> 0 \<Longrightarrow> 0 < y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
  1154   unfolding rapprox_rat_def round_up_def
  1155   by (auto simp: field_simps mult_le_0_iff)
  1156 
  1157 lemma real_divl: "real_divl prec x y \<le> x / y"
  1158   by (simp add: real_divl_def round_down)
  1159 
  1160 lemma real_divr: "x / y \<le> real_divr prec x y"
  1161   using round_up by (simp add: real_divr_def)
  1162 
  1163 lemma float_divl: "real (float_divl prec x y) \<le> real x / real y"
  1164   by transfer (rule real_divl)
  1165 
  1166 lemma real_divl_lower_bound:
  1167   "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_divl prec x y"
  1168   by (simp add: real_divl_def round_down_def zero_le_mult_iff zero_le_divide_iff)
  1169 
  1170 lemma float_divl_lower_bound:
  1171   "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real (float_divl prec x y)"
  1172   by transfer (rule real_divl_lower_bound)
  1173 
  1174 lemma exponent_1: "exponent 1 = 0"
  1175   using exponent_float[of 1 0] by (simp add: one_float_def)
  1176 
  1177 lemma mantissa_1: "mantissa 1 = 1"
  1178   using mantissa_float[of 1 0] by (simp add: one_float_def)
  1179 
  1180 lemma bitlen_1: "bitlen 1 = 1"
  1181   by (simp add: bitlen_def)
  1182 
  1183 lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0"
  1184 proof
  1185   assume "mantissa x = 0" hence z: "0 = real x" using mantissa_exponent by simp
  1186   show "x = 0" by (simp add: zero_float_def z)
  1187 qed (simp add: zero_float_def)
  1188 
  1189 lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)"
  1190 proof (cases "x = 0", simp)
  1191   assume "x \<noteq> 0" hence "mantissa x \<noteq> 0" using mantissa_eq_zero_iff by auto
  1192   have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent)
  1193   also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp
  1194   also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
  1195     using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg `mantissa x \<noteq> 0`
  1196     by (simp add: powr_int) (simp only: two_real_int int_of_reals real_of_int_abs[symmetric]
  1197       real_of_int_le_iff less_imp_le)
  1198   finally show ?thesis by (simp add: powr_add)
  1199 qed
  1200 
  1201 lemma real_divl_pos_less1_bound:
  1202   "0 < x \<Longrightarrow> x < 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real_divl prec 1 x"
  1203 proof (unfold real_divl_def)
  1204   fix prec :: nat and x :: real assume x: "0 < x" "x < 1" and prec: "1 \<le> prec"
  1205   def p \<equiv> "int prec + \<lfloor>log 2 \<bar>x\<bar>\<rfloor>"
  1206   show "1 \<le> round_down (int prec + \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - \<lfloor>log 2 \<bar>1\<bar>\<rfloor>) (1 / x) "
  1207   proof cases
  1208     assume nonneg: "0 \<le> p"
  1209     hence "2 powr real (p) = floor (real ((2::int) ^ nat p)) * floor (1::real)"
  1210       by (simp add: powr_int del: real_of_int_power) simp
  1211     also have "floor (1::real) \<le> floor (1 / x)" using x prec by simp
  1212     also have "floor (real ((2::int) ^ nat p)) * floor (1 / x) \<le>
  1213       floor (real ((2::int) ^ nat p) * (1 / x))"
  1214       by (rule le_mult_floor) (auto simp: x prec less_imp_le)
  1215     finally have "2 powr real p \<le> floor (2 powr nat p / x)" by (simp add: powr_realpow)
  1216     thus ?thesis unfolding p_def[symmetric]
  1217       using x prec nonneg by (simp add: powr_minus inverse_eq_divide round_down_def)
  1218   next
  1219     assume neg: "\<not> 0 \<le> p"
  1220 
  1221     have "x = 2 powr (log 2 x)"
  1222       using x by simp
  1223     also have "2 powr (log 2 x) \<le> 2 powr p"
  1224     proof (rule powr_mono)
  1225       have "log 2 x \<le> \<lceil>log 2 x\<rceil>"
  1226         by simp
  1227       also have "\<dots> \<le> \<lfloor>log 2 x\<rfloor> + 1"
  1228         using ceiling_diff_floor_le_1[of "log 2 x"] by simp
  1229       also have "\<dots> \<le> \<lfloor>log 2 x\<rfloor> + prec"
  1230         using prec by simp
  1231       finally show "log 2 x \<le> real p"
  1232         using x by (simp add: p_def)
  1233     qed simp
  1234     finally have x_le: "x \<le> 2 powr p" .
  1235 
  1236     from neg have "2 powr real p \<le> 2 powr 0"
  1237       by (intro powr_mono) auto
  1238     also have "\<dots> \<le> \<lfloor>2 powr 0\<rfloor>" by simp
  1239     also have "\<dots> \<le> \<lfloor>2 powr real p / x\<rfloor>" unfolding real_of_int_le_iff
  1240       using x x_le by (intro floor_mono) (simp add:  pos_le_divide_eq)
  1241     finally show ?thesis
  1242       using prec x unfolding p_def[symmetric]
  1243       by (simp add: round_down_def powr_minus_divide pos_le_divide_eq)
  1244   qed
  1245 qed
  1246 
  1247 lemma float_divl_pos_less1_bound:
  1248   "0 < real x \<Longrightarrow> real x < 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real (float_divl prec 1 x)"
  1249   by (transfer, rule real_divl_pos_less1_bound)
  1250 
  1251 lemma float_divr: "real x / real y \<le> real (float_divr prec x y)"
  1252   by transfer (rule real_divr)
  1253 
  1254 lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x < 1" shows "1 \<le> real_divr prec 1 x"
  1255 proof -
  1256   have "1 \<le> 1 / x" using `0 < x` and `x < 1` by auto
  1257   also have "\<dots> \<le> real_divr prec 1 x" using real_divr[where x=1 and y=x] by auto
  1258   finally show ?thesis by auto
  1259 qed
  1260 
  1261 lemma float_divr_pos_less1_lower_bound: "0 < x \<Longrightarrow> x < 1 \<Longrightarrow> 1 \<le> float_divr prec 1 x"
  1262   by transfer (rule real_divr_pos_less1_lower_bound)
  1263 
  1264 lemma real_divr_nonpos_pos_upper_bound:
  1265   "x \<le> 0 \<Longrightarrow> 0 < y \<Longrightarrow> real_divr prec x y \<le> 0"
  1266   by (auto simp: field_simps mult_le_0_iff divide_le_0_iff round_up_def real_divr_def)
  1267 
  1268 lemma float_divr_nonpos_pos_upper_bound:
  1269   "real x \<le> 0 \<Longrightarrow> 0 < real y \<Longrightarrow> real (float_divr prec x y) \<le> 0"
  1270   by transfer (rule real_divr_nonpos_pos_upper_bound)
  1271 
  1272 lemma real_divr_nonneg_neg_upper_bound:
  1273   "0 \<le> x \<Longrightarrow> y < 0 \<Longrightarrow> real_divr prec x y \<le> 0"
  1274   by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff divide_le_0_iff round_up_def real_divr_def)
  1275 
  1276 lemma float_divr_nonneg_neg_upper_bound:
  1277   "0 \<le> real x \<Longrightarrow> real y < 0 \<Longrightarrow> real (float_divr prec x y) \<le> 0"
  1278   by transfer (rule real_divr_nonneg_neg_upper_bound)
  1279 
  1280 definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real" where
  1281   "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
  1282 
  1283 lemma truncate_down: "truncate_down prec x \<le> x"
  1284   using round_down by (simp add: truncate_down_def)
  1285 
  1286 lemma truncate_down_le: "x \<le> y \<Longrightarrow> truncate_down prec x \<le> y"
  1287   by (rule order_trans[OF truncate_down])
  1288 
  1289 definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real" where
  1290   "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
  1291 
  1292 lemma truncate_up: "x \<le> truncate_up prec x"
  1293   using round_up by (simp add: truncate_up_def)
  1294 
  1295 lemma truncate_up_le: "x \<le> y \<Longrightarrow> x \<le> truncate_up prec y"
  1296   by (rule order_trans[OF _ truncate_up])
  1297 
  1298 lemma truncate_up_zero[simp]: "truncate_up prec 0 = 0"
  1299   by (simp add: truncate_up_def)
  1300 
  1301 lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_up
  1302   by (simp add: truncate_up_def)
  1303 
  1304 lemma float_round_up: "real x \<le> real (float_round_up prec x)"
  1305   using truncate_up by transfer simp
  1306 
  1307 lift_definition float_round_down :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_down
  1308   by (simp add: truncate_down_def)
  1309 
  1310 lemma float_round_down: "real (float_round_down prec x) \<le> real x"
  1311   using truncate_down by transfer simp
  1312 
  1313 lemma floor_add2[simp]: "\<lfloor> real i + x \<rfloor> = i + \<lfloor> x \<rfloor>"
  1314   using floor_add[of x i] by (simp del: floor_add add: ac_simps)
  1315 
  1316 lemma compute_float_round_down[code]:
  1317   "float_round_down prec (Float m e) = (let d = bitlen (abs m) - int prec in
  1318     if 0 < d then let P = 2^nat d ; n = m div P in Float n (e + d)
  1319              else Float m e)"
  1320   using Float.compute_float_down[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric]
  1321   by transfer (simp add: field_simps abs_mult log_mult bitlen_def truncate_down_def
  1322     cong del: if_weak_cong)
  1323 hide_fact (open) compute_float_round_down
  1324 
  1325 lemma compute_float_round_up[code]:
  1326   "float_round_up prec (Float m e) = (let d = (bitlen (abs m) - int prec) in
  1327      if 0 < d then let P = 2^nat d ; n = m div P ; r = m mod P
  1328                    in Float (n + (if r = 0 then 0 else 1)) (e + d)
  1329               else Float m e)"
  1330   using Float.compute_float_up[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric]
  1331   unfolding Let_def
  1332   by transfer (simp add: field_simps abs_mult log_mult bitlen_def truncate_up_def
  1333     cong del: if_weak_cong)
  1334 hide_fact (open) compute_float_round_up
  1335 
  1336 lemma round_up_mono: "x \<le> y \<Longrightarrow> round_up p x \<le> round_up p y"
  1337   by (auto intro!: ceiling_mono simp: round_up_def)
  1338 
  1339 lemma truncate_up_nonneg_mono:
  1340   assumes "0 \<le> x" "x \<le> y"
  1341   shows "truncate_up prec x \<le> truncate_up prec y"
  1342 proof -
  1343   {
  1344     assume "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>"
  1345     hence ?thesis
  1346       using assms
  1347       by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono)
  1348   } moreover {
  1349     assume "0 < x"
  1350     hence "log 2 x \<le> log 2 y" using assms by auto
  1351     moreover
  1352     assume "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>"
  1353     ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
  1354       unfolding atomize_conj
  1355       by (metis floor_less_cancel linorder_cases not_le)
  1356     have "truncate_up prec x =
  1357       real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> * 2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1)"
  1358       using assms by (simp add: truncate_up_def round_up_def)
  1359     also have "\<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> (2 ^ prec)"
  1360     proof (unfold ceiling_le_eq)
  1361       have "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> x * (2 powr real prec / (2 powr log 2 x))"
  1362         using real_of_int_floor_add_one_ge[of "log 2 x"] assms
  1363         by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono)
  1364       thus "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> real ((2::int) ^ prec)"
  1365         using `0 < x` by (simp add: powr_realpow)
  1366     qed
  1367     hence "real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> 2 powr int prec"
  1368       by (auto simp: powr_realpow)
  1369     also
  1370     have "2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)"
  1371       using logless flogless by (auto intro!: floor_mono)
  1372     also have "2 powr real (int prec) \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>))"
  1373       using assms `0 < x`
  1374       by (auto simp: algebra_simps)
  1375     finally have "truncate_up prec x \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>)) * 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)"
  1376       by simp
  1377     also have "\<dots> = 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>) - real (int prec - \<lfloor>log 2 y\<rfloor>))"
  1378       by (subst powr_add[symmetric]) simp
  1379     also have "\<dots> = y"
  1380       using `0 < x` assms
  1381       by (simp add: powr_add)
  1382     also have "\<dots> \<le> truncate_up prec y"
  1383       by (rule truncate_up)
  1384     finally have ?thesis .
  1385   } moreover {
  1386     assume "~ 0 < x"
  1387     hence ?thesis
  1388       using assms
  1389       by (auto intro!: truncate_up_le)
  1390   } ultimately show ?thesis
  1391     by blast
  1392 qed
  1393 
  1394 lemma truncate_up_nonpos: "x \<le> 0 \<Longrightarrow> truncate_up prec x \<le> 0"
  1395   by (auto simp: truncate_up_def round_up_def intro!: mult_nonpos_nonneg)
  1396 
  1397 lemma truncate_down_nonpos: "x \<le> 0 \<Longrightarrow> truncate_down prec x \<le> 0"
  1398   by (auto simp: truncate_down_def round_down_def intro!: mult_nonpos_nonneg
  1399     order_le_less_trans[of _ 0, OF mult_nonpos_nonneg])
  1400 
  1401 lemma truncate_up_switch_sign_mono:
  1402   assumes "x \<le> 0" "0 \<le> y"
  1403   shows "truncate_up prec x \<le> truncate_up prec y"
  1404 proof -
  1405   note truncate_up_nonpos[OF `x \<le> 0`]
  1406   also note truncate_up_le[OF `0 \<le> y`]
  1407   finally show ?thesis .
  1408 qed
  1409 
  1410 lemma truncate_down_zeroprec_mono:
  1411   assumes "0 < x" "x \<le> y"
  1412   shows "truncate_down 0 x \<le> truncate_down 0 y"
  1413 proof -
  1414   have "x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1) = x * inverse (2 powr ((real \<lfloor>log 2 x\<rfloor> + 1)))"
  1415     by (simp add: powr_divide2[symmetric] powr_add powr_minus inverse_eq_divide)
  1416   also have "\<dots> = 2 powr (log 2 x - (real \<lfloor>log 2 x\<rfloor>) - 1)"
  1417     using `0 < x`
  1418     by (auto simp: inverse_eq_divide field_simps powr_add powr_divide2[symmetric])
  1419   also have "\<dots> < 2 powr 0"
  1420     using real_of_int_floor_add_one_gt
  1421     unfolding neg_less_iff_less
  1422     by (intro powr_less_mono) (auto simp: algebra_simps)
  1423   finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> < 1"
  1424     unfolding less_ceiling_eq real_of_int_minus real_of_one
  1425     by simp
  1426   moreover
  1427   have "0 \<le> \<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor>"
  1428     using `x > 0` by auto
  1429   ultimately have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> \<in> {0 ..< 1}"
  1430     by simp
  1431   also have "\<dots> \<subseteq> {0}" by auto
  1432   finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> = 0" by simp
  1433   with assms show ?thesis
  1434     by (auto simp: truncate_down_def round_down_def)
  1435 qed
  1436 
  1437 lemma truncate_down_nonneg: "0 \<le> y \<Longrightarrow> 0 \<le> truncate_down prec y"
  1438   by (auto simp: truncate_down_def round_down_def)
  1439 
  1440 lemma truncate_down_zero: "truncate_down prec 0 = 0"
  1441   by (auto simp: truncate_down_def round_down_def)
  1442 
  1443 lemma truncate_down_switch_sign_mono:
  1444   assumes "x \<le> 0" "0 \<le> y"
  1445   assumes "x \<le> y"
  1446   shows "truncate_down prec x \<le> truncate_down prec y"
  1447 proof -
  1448   note truncate_down_nonpos[OF `x \<le> 0`]
  1449   also note truncate_down_nonneg[OF `0 \<le> y`]
  1450   finally show ?thesis .
  1451 qed
  1452 
  1453 lemma truncate_up_uminus_truncate_down:
  1454   "truncate_up prec x = - truncate_down prec (- x)"
  1455   "truncate_up prec (-x) = - truncate_down prec x"
  1456   by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def)
  1457 
  1458 lemma truncate_down_uminus_truncate_up:
  1459   "truncate_down prec x = - truncate_up prec (- x)"
  1460   "truncate_down prec (-x) = - truncate_up prec x"
  1461   by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def)
  1462 
  1463 lemma truncate_down_nonneg_mono:
  1464   assumes "0 \<le> x" "x \<le> y"
  1465   shows "truncate_down prec x \<le> truncate_down prec y"
  1466 proof -
  1467   {
  1468     assume "0 < x" "prec = 0"
  1469     with assms have ?thesis
  1470       by (simp add: truncate_down_zeroprec_mono)
  1471   } moreover {
  1472     assume "~ 0 < x"
  1473     with assms have "x = 0" "0 \<le> y" by simp_all
  1474     hence ?thesis
  1475       by (auto simp add: truncate_down_zero intro!: truncate_down_nonneg)
  1476   } moreover {
  1477     assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
  1478     hence ?thesis
  1479       using assms
  1480       by (auto simp: truncate_down_def round_down_def intro!: floor_mono)
  1481   } moreover {
  1482     assume "0 < x"
  1483     hence "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y" using assms by auto
  1484     moreover
  1485     assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
  1486     ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
  1487       unfolding atomize_conj abs_of_pos[OF `0 < x`] abs_of_pos[OF `0 < y`]
  1488       by (metis floor_less_cancel linorder_cases not_le)
  1489     assume "prec \<noteq> 0" hence [simp]: "prec \<ge> Suc 0" by auto
  1490     have "2 powr (prec - 1) \<le> y * 2 powr real (prec - 1) / (2 powr log 2 y)"
  1491       using `0 < y`
  1492       by simp
  1493     also have "\<dots> \<le> y * 2 powr real prec / (2 powr (real \<lfloor>log 2 y\<rfloor> + 1))"
  1494       using `0 \<le> y` `0 \<le> x` assms(2)
  1495       by (auto intro!: powr_mono divide_left_mono
  1496         simp: real_of_nat_diff powr_add
  1497         powr_divide2[symmetric])
  1498     also have "\<dots> = y * 2 powr real prec / (2 powr real \<lfloor>log 2 y\<rfloor> * 2)"
  1499       by (auto simp: powr_add)
  1500     finally have "(2 ^ (prec - 1)) \<le> \<lfloor>y * 2 powr real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
  1501       using `0 \<le> y`
  1502       by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow)
  1503     hence "(2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1) \<le> truncate_down prec y"
  1504       by (auto simp: truncate_down_def round_down_def)
  1505     moreover
  1506     {
  1507       have "x = 2 powr (log 2 \<bar>x\<bar>)" using `0 < x` by simp
  1508       also have "\<dots> \<le> (2 ^ (prec )) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)"
  1509         using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"]
  1510         by (auto simp: powr_realpow[symmetric] powr_add[symmetric] algebra_simps)
  1511       also
  1512       have "2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)"
  1513         using logless flogless `x > 0` `y > 0`
  1514         by (auto intro!: floor_mono)
  1515       finally have "x \<le> (2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)"
  1516         by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms real_of_nat_diff)
  1517     } ultimately have ?thesis
  1518       by (metis dual_order.trans truncate_down)
  1519   } ultimately show ?thesis by blast
  1520 qed
  1521 
  1522 lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y"
  1523   apply (cases "0 \<le> x")
  1524   apply (rule truncate_down_nonneg_mono, assumption+)
  1525   apply (simp add: truncate_down_uminus_truncate_up)
  1526   apply (cases "0 \<le> y")
  1527   apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono)
  1528   done
  1529 
  1530 lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y"
  1531   by (simp add: truncate_up_uminus_truncate_down truncate_down_mono)
  1532 
  1533 lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0"
  1534  apply (auto simp: zero_float_def mult_le_0_iff)
  1535  using powr_gt_zero[of 2 b] by simp
  1536 
  1537 lemma real_of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)"
  1538   unfolding pprt_def sup_float_def max_def sup_real_def by auto
  1539 
  1540 lemma real_of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)"
  1541   unfolding nprt_def inf_float_def min_def inf_real_def by auto
  1542 
  1543 lift_definition int_floor_fl :: "float \<Rightarrow> int" is floor .
  1544 
  1545 lemma compute_int_floor_fl[code]:
  1546   "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))"
  1547   by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
  1548 hide_fact (open) compute_int_floor_fl
  1549 
  1550 lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real (floor x)" by simp
  1551 
  1552 lemma compute_floor_fl[code]:
  1553   "floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)"
  1554   by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
  1555 hide_fact (open) compute_floor_fl
  1556 
  1557 lemma floor_fl: "real (floor_fl x) \<le> real x" by transfer simp
  1558 
  1559 lemma int_floor_fl: "real (int_floor_fl x) \<le> real x" by transfer simp
  1560 
  1561 lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0"
  1562 proof (cases "floor_fl x = float_of 0")
  1563   case True
  1564   then show ?thesis by (simp add: floor_fl_def)
  1565 next
  1566   case False
  1567   have eq: "floor_fl x = Float \<lfloor>real x\<rfloor> 0" by transfer simp
  1568   obtain i where "\<lfloor>real x\<rfloor> = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i"
  1569     by (rule denormalize_shift[OF eq[THEN eq_reflection] False])
  1570   then show ?thesis by simp
  1571 qed
  1572 
  1573 end
  1574