src/HOL/Library/Float.thy
author wenzelm
Sat, 23 May 2015 17:19:37 +0200
changeset 60299 5ae2a2e74c93
parent 60017 b785d6d06430
child 60376 528a48f4ad87
permissions -rw-r--r--
clarified NEWS: document_files are officially required since Isabelle2014, but the absence was tolerated as legacy feature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47615
341fd902ef1c transfer now handles Let
hoelzl
parents: 47608
diff changeset
     1
(*  Title:      HOL/Library/Float.thy
341fd902ef1c transfer now handles Let
hoelzl
parents: 47608
diff changeset
     2
    Author:     Johannes Hölzl, Fabian Immler
341fd902ef1c transfer now handles Let
hoelzl
parents: 47608
diff changeset
     3
    Copyright   2012  TU München
341fd902ef1c transfer now handles Let
hoelzl
parents: 47608
diff changeset
     4
*)
341fd902ef1c transfer now handles Let
hoelzl
parents: 47608
diff changeset
     5
58881
b9556a055632 modernized header;
wenzelm
parents: 58834
diff changeset
     6
section {* Floating-Point Numbers *}
29988
747f0c519090 add header
huffman
parents: 29804
diff changeset
     7
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
     8
theory Float
51542
738598beeb26 tuned imports;
wenzelm
parents: 51375
diff changeset
     9
imports Complex_Main Lattice_Algebras
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
    10
begin
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    11
49812
e3945ddcb9aa eliminated some remaining uses of typedef with implicit set definition;
wenzelm
parents: 47937
diff changeset
    12
definition "float = {m * 2 powr e | (m :: int) (e :: int). True}"
e3945ddcb9aa eliminated some remaining uses of typedef with implicit set definition;
wenzelm
parents: 47937
diff changeset
    13
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 49812
diff changeset
    14
typedef float = float
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    15
  morphisms real_of_float float_of
49812
e3945ddcb9aa eliminated some remaining uses of typedef with implicit set definition;
wenzelm
parents: 47937
diff changeset
    16
  unfolding float_def by auto
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    17
58042
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    18
instantiation float :: real_of
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    19
begin
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    20
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    21
definition real_float :: "float \<Rightarrow> real" where
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
    22
  real_of_float_def[code_unfold]: "real \<equiv> real_of_float"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
    23
58042
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    24
instance ..
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    25
end
ffa9e39763e3 introduce real_of typeclass for real :: 'a => real
hoelzl
parents: 57862
diff changeset
    26
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
    27
lemma type_definition_float': "type_definition real float_of float"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
    28
  using type_definition_float unfolding real_of_float_def .
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
    29
59487
adaa430fc0f7 default abstypes and default abstract equations make technical (no_code) annotation superfluous
haftmann
parents: 58989
diff changeset
    30
setup_lifting type_definition_float'
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    31
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    32
lemmas float_of_inject[simp]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    33
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
    34
declare [[coercion "real :: float \<Rightarrow> real"]]
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
    35
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
    36
lemma real_of_float_eq:
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
    37
  fixes f1 f2 :: float shows "f1 = f2 \<longleftrightarrow> real f1 = real f2"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    38
  unfolding real_of_float_def real_of_float_inject ..
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    39
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    40
lemma float_of_real[simp]: "float_of (real x) = x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    41
  unfolding real_of_float_def by (rule real_of_float_inverse)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
    42
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    43
lemma real_float[simp]: "x \<in> float \<Longrightarrow> real (float_of x) = x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    44
  unfolding real_of_float_def by (rule float_of_inverse)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    45
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    46
subsection {* Real operations preserving the representation as floating point number *}
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    47
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    48
lemma floatI: fixes m e :: int shows "m * 2 powr e = x \<Longrightarrow> x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    49
  by (auto simp: float_def)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    50
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    51
lemma zero_float[simp]: "0 \<in> float" by (auto simp: float_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    52
lemma one_float[simp]: "1 \<in> float" by (intro floatI[of 1 0]) simp
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
    53
lemma numeral_float[simp]: "numeral i \<in> float" by (intro floatI[of "numeral i" 0]) simp
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
    54
lemma neg_numeral_float[simp]: "- numeral i \<in> float" by (intro floatI[of "- numeral i" 0]) simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    55
lemma real_of_int_float[simp]: "real (x :: int) \<in> float" by (intro floatI[of x 0]) simp
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
    56
lemma real_of_nat_float[simp]: "real (x :: nat) \<in> float" by (intro floatI[of x 0]) simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    57
lemma two_powr_int_float[simp]: "2 powr (real (i::int)) \<in> float" by (intro floatI[of 1 i]) simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    58
lemma two_powr_nat_float[simp]: "2 powr (real (i::nat)) \<in> float" by (intro floatI[of 1 i]) simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    59
lemma two_powr_minus_int_float[simp]: "2 powr - (real (i::int)) \<in> float" by (intro floatI[of 1 "-i"]) simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    60
lemma two_powr_minus_nat_float[simp]: "2 powr - (real (i::nat)) \<in> float" by (intro floatI[of 1 "-i"]) simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    61
lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float" by (intro floatI[of 1 "numeral i"]) simp
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
    62
lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float" by (intro floatI[of 1 "- numeral i"]) simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    63
lemma two_pow_float[simp]: "2 ^ n \<in> float" by (intro floatI[of 1 "n"]) (simp add: powr_realpow)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    64
lemma real_of_float_float[simp]: "real (f::float) \<in> float" by (cases f) simp
45495
c55a07526dbe cleaned up float theories; removed duplicate definitions and theorems
hoelzl
parents: 44766
diff changeset
    65
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    66
lemma plus_float[simp]: "r \<in> float \<Longrightarrow> p \<in> float \<Longrightarrow> r + p \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    67
  unfolding float_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    68
proof (safe, simp)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    69
  fix e1 m1 e2 m2 :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    70
  { fix e1 m1 e2 m2 :: int assume "e1 \<le> e2"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    71
    then have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    72
      by (simp add: powr_realpow[symmetric] powr_divide2[symmetric] field_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    73
    then have "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    74
      by blast }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    75
  note * = this
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    76
  show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    77
  proof (cases e1 e2 rule: linorder_le_cases)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    78
    assume "e2 \<le> e1" from *[OF this, of m2 m1] show ?thesis by (simp add: ac_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    79
  qed (rule *)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    80
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    81
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    82
lemma uminus_float[simp]: "x \<in> float \<Longrightarrow> -x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    83
  apply (auto simp: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
    84
  apply hypsubst_thin
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
    85
  apply (rename_tac m e)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
    86
  apply (rule_tac x="-m" in exI)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
    87
  apply (rule_tac x="e" in exI)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    88
  apply (simp add: field_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    89
  done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
    90
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    91
lemma times_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x * y \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    92
  apply (auto simp: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
    93
  apply hypsubst_thin
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
    94
  apply (rename_tac mx my ex ey)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
    95
  apply (rule_tac x="mx * my" in exI)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
    96
  apply (rule_tac x="ex + ey" in exI)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    97
  apply (simp add: powr_add)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    98
  done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
    99
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   100
lemma minus_float[simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> x - y \<in> float"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53381
diff changeset
   101
  using plus_float [of x "- y"] by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   102
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   103
lemma abs_float[simp]: "x \<in> float \<Longrightarrow> abs x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   104
  by (cases x rule: linorder_cases[of 0]) auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   105
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   106
lemma sgn_of_float[simp]: "x \<in> float \<Longrightarrow> sgn x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   107
  by (cases x rule: linorder_cases[of 0]) (auto intro!: uminus_float)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
   108
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   109
lemma div_power_2_float[simp]: "x \<in> float \<Longrightarrow> x / 2^d \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   110
  apply (auto simp add: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
   111
  apply hypsubst_thin
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   112
  apply (rename_tac m e)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   113
  apply (rule_tac x="m" in exI)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   114
  apply (rule_tac x="e - d" in exI)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   115
  apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   116
  done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   117
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   118
lemma div_power_2_int_float[simp]: "x \<in> float \<Longrightarrow> x / (2::int)^d \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   119
  apply (auto simp add: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
   120
  apply hypsubst_thin
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   121
  apply (rename_tac m e)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   122
  apply (rule_tac x="m" in exI)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   123
  apply (rule_tac x="e - d" in exI)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   124
  apply (simp add: powr_realpow[symmetric] field_simps powr_add[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   125
  done
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   126
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   127
lemma div_numeral_Bit0_float[simp]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   128
  assumes x: "x / numeral n \<in> float" shows "x / (numeral (Num.Bit0 n)) \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   129
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   130
  have "(x / numeral n) / 2^1 \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   131
    by (intro x div_power_2_float)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   132
  also have "(x / numeral n) / 2^1 = x / (numeral (Num.Bit0 n))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   133
    by (induct n) auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   134
  finally show ?thesis .
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   135
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   136
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   137
lemma div_neg_numeral_Bit0_float[simp]:
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   138
  assumes x: "x / numeral n \<in> float" shows "x / (- numeral (Num.Bit0 n)) \<in> float"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   139
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   140
  have "- (x / numeral (Num.Bit0 n)) \<in> float" using x by simp
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   141
  also have "- (x / numeral (Num.Bit0 n)) = x / - numeral (Num.Bit0 n)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   142
    by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   143
  finally show ?thesis .
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   144
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   145
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   146
lemma power_float[simp]: assumes "a \<in> float" shows "a ^ b \<in> float"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   147
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   148
  from assms obtain m e::int where "a = m * 2 powr e"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   149
    by (auto simp: float_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   150
  thus ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   151
    by (auto intro!: floatI[where m="m^b" and e = "e*b"]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   152
      simp: power_mult_distrib powr_realpow[symmetric] powr_powr)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   153
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   154
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   155
lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e" by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   156
declare Float.rep_eq[simp]
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   157
47780
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   158
lemma compute_real_of_float[code]:
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   159
  "real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))"
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   160
by (simp add: real_of_float_def[symmetric] powr_int)
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   161
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   162
code_datatype Float
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   163
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   164
subsection {* Arithmetic operations on floating point numbers *}
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   165
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   166
instantiation float :: "{ring_1, linorder, linordered_ring, linordered_idom, numeral, equal}"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   167
begin
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   168
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   169
lift_definition zero_float :: float is 0 by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   170
declare zero_float.rep_eq[simp]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   171
lift_definition one_float :: float is 1 by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   172
declare one_float.rep_eq[simp]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   173
lift_definition plus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op +" by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   174
declare plus_float.rep_eq[simp]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   175
lift_definition times_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op *" by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   176
declare times_float.rep_eq[simp]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   177
lift_definition minus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "op -" by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   178
declare minus_float.rep_eq[simp]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   179
lift_definition uminus_float :: "float \<Rightarrow> float" is "uminus" by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   180
declare uminus_float.rep_eq[simp]
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   181
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   182
lift_definition abs_float :: "float \<Rightarrow> float" is abs by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   183
declare abs_float.rep_eq[simp]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   184
lift_definition sgn_float :: "float \<Rightarrow> float" is sgn by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   185
declare sgn_float.rep_eq[simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   186
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
   187
lift_definition equal_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op = :: real \<Rightarrow> real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   188
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
   189
lift_definition less_eq_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op \<le>" .
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   190
declare less_eq_float.rep_eq[simp]
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
   191
lift_definition less_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "op <" .
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   192
declare less_float.rep_eq[simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   193
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   194
instance
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   195
  proof qed (transfer, fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   196
end
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   197
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   198
lemma Float_0_eq_0[simp]: "Float 0 e = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   199
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   200
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   201
lemma real_of_float_power[simp]: fixes f::float shows "real (f^n) = real f^n"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   202
  by (induct n) simp_all
45495
c55a07526dbe cleaned up float theories; removed duplicate definitions and theorems
hoelzl
parents: 44766
diff changeset
   203
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   204
lemma fixes x y::float
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   205
  shows real_of_float_min: "real (min x y) = min (real x) (real y)"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   206
    and real_of_float_max: "real (max x y) = max (real x) (real y)"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   207
  by (simp_all add: min_def max_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   208
53215
5e47c31c6f7c renamed typeclass dense_linorder to unbounded_dense_linorder
hoelzl
parents: 51542
diff changeset
   209
instance float :: unbounded_dense_linorder
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   210
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   211
  fix a b :: float
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   212
  show "\<exists>c. a < c"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   213
    apply (intro exI[of _ "a + 1"])
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   214
    apply transfer
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   215
    apply simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   216
    done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   217
  show "\<exists>c. c < a"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   218
    apply (intro exI[of _ "a - 1"])
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   219
    apply transfer
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   220
    apply simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   221
    done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   222
  assume "a < b"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   223
  then show "\<exists>c. a < c \<and> c < b"
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58042
diff changeset
   224
    apply (intro exI[of _ "(a + b) * Float 1 (- 1)"])
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   225
    apply transfer
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   226
    apply (simp add: powr_minus)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   227
    done
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   228
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   229
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   230
instantiation float :: lattice_ab_group_add
46573
8c4c5c8dcf7a misc tuning;
wenzelm
parents: 46028
diff changeset
   231
begin
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   232
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   233
definition inf_float::"float\<Rightarrow>float\<Rightarrow>float"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   234
where "inf_float a b = min a b"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   235
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   236
definition sup_float::"float\<Rightarrow>float\<Rightarrow>float"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   237
where "sup_float a b = max a b"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   238
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   239
instance
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   240
  by default
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   241
     (transfer, simp_all add: inf_float_def sup_float_def real_of_float_min real_of_float_max)+
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   242
end
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   243
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   244
lemma float_numeral[simp]: "real (numeral x :: float) = numeral x"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   245
  apply (induct x)
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   246
  apply simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   247
  apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq real_float
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   248
                  plus_float.rep_eq one_float.rep_eq plus_float numeral_float one_float)
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   249
  done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   250
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   251
lemma transfer_numeral [transfer_rule]:
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55565
diff changeset
   252
  "rel_fun (op =) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55565
diff changeset
   253
  unfolding rel_fun_def float.pcr_cr_eq  cr_float_def by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   254
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   255
lemma float_neg_numeral[simp]: "real (- numeral x :: float) = - numeral x"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   256
  by simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   257
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   258
lemma transfer_neg_numeral [transfer_rule]:
55945
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55565
diff changeset
   259
  "rel_fun (op =) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)"
e96383acecf9 renamed 'fun_rel' to 'rel_fun'
blanchet
parents: 55565
diff changeset
   260
  unfolding rel_fun_def float.pcr_cr_eq cr_float_def by simp
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   261
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   262
lemma
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   263
  shows float_of_numeral[simp]: "numeral k = float_of (numeral k)"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   264
    and float_of_neg_numeral[simp]: "- numeral k = float_of (- numeral k)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   265
  unfolding real_of_float_eq by simp_all
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   266
58987
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   267
subsection {* Quickcheck *}
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   268
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   269
instantiation float :: exhaustive
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   270
begin
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   271
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   272
definition exhaustive_float where
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   273
  "exhaustive_float f d =
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   274
    Quickcheck_Exhaustive.exhaustive (%x. Quickcheck_Exhaustive.exhaustive (%y. f (Float x y)) d) d"
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   275
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   276
instance ..
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   277
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   278
end
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   279
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   280
definition (in term_syntax) [code_unfold]:
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   281
  "valtermify_float x y = Code_Evaluation.valtermify Float {\<cdot>} x {\<cdot>} y"
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   282
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   283
instantiation float :: full_exhaustive
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   284
begin
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   285
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   286
definition full_exhaustive_float where
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   287
  "full_exhaustive_float f d =
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   288
    Quickcheck_Exhaustive.full_exhaustive
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   289
      (\<lambda>x. Quickcheck_Exhaustive.full_exhaustive (\<lambda>y. f (valtermify_float x y)) d) d"
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   290
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   291
instance ..
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   292
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   293
end
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   294
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   295
instantiation float :: random
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   296
begin
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   297
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   298
definition "Quickcheck_Random.random i =
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   299
  scomp (Quickcheck_Random.random (2 ^ nat_of_natural i))
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   300
    (\<lambda>man. scomp (Quickcheck_Random.random i) (\<lambda>exp. Pair (valtermify_float man exp)))"
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   301
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   302
instance ..
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   303
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   304
end
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   305
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   306
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   307
subsection {* Represent floats as unique mantissa and exponent *}
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   308
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   309
lemma int_induct_abs[case_names less]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   310
  fixes j :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   311
  assumes H: "\<And>n. (\<And>i. \<bar>i\<bar> < \<bar>n\<bar> \<Longrightarrow> P i) \<Longrightarrow> P n"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   312
  shows "P j"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   313
proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   314
  case less show ?case by (rule H[OF less]) simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   315
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   316
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   317
lemma int_cancel_factors:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   318
  fixes n :: int assumes "1 < r" shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   319
proof (induct n rule: int_induct_abs)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   320
  case (less n)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   321
  { fix m assume n: "n \<noteq> 0" "n = m * r"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   322
    then have "\<bar>m \<bar> < \<bar>n\<bar>"
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59487
diff changeset
   323
      using `1 < r` by (simp add: abs_mult)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   324
    from less[OF this] n have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" by auto }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   325
  then show ?case
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59487
diff changeset
   326
    by (metis dvd_def monoid_mult_class.mult.right_neutral mult.commute power_0)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   327
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   328
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   329
lemma mult_powr_eq_mult_powr_iff_asym:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   330
  fixes m1 m2 e1 e2 :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   331
  assumes m1: "\<not> 2 dvd m1" and "e1 \<le> e2"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   332
  shows "m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   333
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   334
  have "m1 \<noteq> 0" using m1 unfolding dvd_def by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   335
  assume eq: "m1 * 2 powr e1 = m2 * 2 powr e2"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   336
  with `e1 \<le> e2` have "m1 = m2 * 2 powr nat (e2 - e1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   337
    by (simp add: powr_divide2[symmetric] field_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   338
  also have "\<dots> = m2 * 2^nat (e2 - e1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   339
    by (simp add: powr_realpow)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   340
  finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   341
    unfolding real_of_int_inject .
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   342
  with m1 have "m1 = m2"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   343
    by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   344
  then show "m1 = m2 \<and> e1 = e2"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   345
    using eq `m1 \<noteq> 0` by (simp add: powr_inj)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   346
qed simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   347
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   348
lemma mult_powr_eq_mult_powr_iff:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   349
  fixes m1 m2 e1 e2 :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   350
  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"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   351
  using mult_powr_eq_mult_powr_iff_asym[of m1 e1 e2 m2]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   352
  using mult_powr_eq_mult_powr_iff_asym[of m2 e2 e1 m1]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   353
  by (cases e1 e2 rule: linorder_le_cases) auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   354
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   355
lemma floatE_normed:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   356
  assumes x: "x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   357
  obtains (zero) "x = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   358
   | (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   359
proof atomize_elim
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   360
  { assume "x \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   361
    from x obtain m e :: int where x: "x = m * 2 powr e" by (auto simp: float_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   362
    with `x \<noteq> 0` int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   363
      by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   364
    with `\<not> 2 dvd k` x have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   365
      by (rule_tac exI[of _ "k"], rule_tac exI[of _ "e + int i"])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   366
         (simp add: powr_add powr_realpow) }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   367
  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)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   368
    by blast
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   369
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   370
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   371
lemma float_normed_cases:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   372
  fixes f :: float
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   373
  obtains (zero) "f = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   374
   | (powr) m e :: int where "real f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   375
proof (atomize_elim, induct f)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   376
  case (float_of y) then show ?case
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   377
    by (cases rule: floatE_normed) (auto simp: zero_float_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   378
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   379
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   380
definition mantissa :: "float \<Rightarrow> int" where
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   381
  "mantissa f = fst (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   382
   \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   383
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   384
definition exponent :: "float \<Rightarrow> int" where
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   385
  "exponent f = snd (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   386
   \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   387
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   388
lemma
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   389
  shows exponent_0[simp]: "exponent (float_of 0) = 0" (is ?E)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   390
    and mantissa_0[simp]: "mantissa (float_of 0) = 0" (is ?M)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   391
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   392
  have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)" by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   393
  then show ?E ?M
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   394
    by (auto simp add: mantissa_def exponent_def zero_float_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   395
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   396
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   397
lemma
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   398
  shows mantissa_exponent: "real f = mantissa f * 2 powr exponent f" (is ?E)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   399
    and mantissa_not_dvd: "f \<noteq> (float_of 0) \<Longrightarrow> \<not> 2 dvd mantissa f" (is "_ \<Longrightarrow> ?D")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   400
proof cases
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   401
  assume [simp]: "f \<noteq> (float_of 0)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   402
  have "f = mantissa f * 2 powr exponent f \<and> \<not> 2 dvd mantissa f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   403
  proof (cases f rule: float_normed_cases)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   404
    case (powr m e)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   405
    then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   406
     \<or> (f \<noteq> 0 \<and> real f = real (fst p) * 2 powr real (snd p) \<and> \<not> 2 dvd fst p)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   407
      by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   408
    then show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   409
      unfolding exponent_def mantissa_def
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   410
      by (rule someI2_ex) (simp add: zero_float_def)
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   411
  qed (simp add: zero_float_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   412
  then show ?E ?D by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   413
qed simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   414
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   415
lemma mantissa_noteq_0: "f \<noteq> float_of 0 \<Longrightarrow> mantissa f \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   416
  using mantissa_not_dvd[of f] by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   417
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   418
lemma
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   419
  fixes m e :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   420
  defines "f \<equiv> float_of (m * 2 powr e)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   421
  assumes dvd: "\<not> 2 dvd m"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   422
  shows mantissa_float: "mantissa f = m" (is "?M")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   423
    and exponent_float: "m \<noteq> 0 \<Longrightarrow> exponent f = e" (is "_ \<Longrightarrow> ?E")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   424
proof cases
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   425
  assume "m = 0" with dvd show "mantissa f = m" by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   426
next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   427
  assume "m \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   428
  then have f_not_0: "f \<noteq> float_of 0" by (simp add: f_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   429
  from mantissa_exponent[of f]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   430
  have "m * 2 powr e = mantissa f * 2 powr exponent f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   431
    by (auto simp add: f_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   432
  then show "?M" "?E"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   433
    using mantissa_not_dvd[OF f_not_0] dvd
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   434
    by (auto simp: mult_powr_eq_mult_powr_iff)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   435
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   436
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   437
subsection {* Compute arithmetic operations *}
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   438
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   439
lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   440
  unfolding real_of_float_eq mantissa_exponent[of f] by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   441
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   442
lemma Float_cases[case_names Float, cases type: float]:
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   443
  fixes f :: float
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   444
  obtains (Float) m e :: int where "f = Float m e"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   445
  using Float_mantissa_exponent[symmetric]
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   446
  by (atomize_elim) auto
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   447
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   448
lemma denormalize_shift:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   449
  assumes f_def: "f \<equiv> Float m e" and not_0: "f \<noteq> float_of 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   450
  obtains i where "m = mantissa f * 2 ^ i" "e = exponent f - i"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   451
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   452
  from mantissa_exponent[of f] f_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   453
  have "m * 2 powr e = mantissa f * 2 powr exponent f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   454
    by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   455
  then have eq: "m = mantissa f * 2 powr (exponent f - e)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   456
    by (simp add: powr_divide2[symmetric] field_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   457
  moreover
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   458
  have "e \<le> exponent f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   459
  proof (rule ccontr)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   460
    assume "\<not> e \<le> exponent f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   461
    then have pos: "exponent f < e" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   462
    then have "2 powr (exponent f - e) = 2 powr - real (e - exponent f)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   463
      by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   464
    also have "\<dots> = 1 / 2^nat (e - exponent f)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   465
      using pos by (simp add: powr_realpow[symmetric] powr_divide2[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   466
    finally have "m * 2^nat (e - exponent f) = real (mantissa f)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   467
      using eq by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   468
    then have "mantissa f = m * 2^nat (e - exponent f)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   469
      unfolding real_of_int_inject by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   470
    with `exponent f < e` have "2 dvd mantissa f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   471
      apply (intro dvdI[where k="m * 2^(nat (e-exponent f)) div 2"])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   472
      apply (cases "nat (e - exponent f)")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   473
      apply auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   474
      done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   475
    then show False using mantissa_not_dvd[OF not_0] by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   476
  qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   477
  ultimately have "real m = mantissa f * 2^nat (exponent f - e)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   478
    by (simp add: powr_realpow[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   479
  with `e \<le> exponent f`
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   480
  show "m = mantissa f * 2 ^ nat (exponent f - e)" "e = exponent f - nat (exponent f - e)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   481
    unfolding real_of_int_inject by auto
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   482
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   483
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   484
lemma compute_float_zero[code_unfold, code]: "0 = Float 0 0"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   485
  by transfer simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   486
hide_fact (open) compute_float_zero
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   487
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   488
lemma compute_float_one[code_unfold, code]: "1 = Float 1 0"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   489
  by transfer simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   490
hide_fact (open) compute_float_one
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   491
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   492
lift_definition normfloat :: "float \<Rightarrow> float" is "\<lambda>x. x" .
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   493
lemma normloat_id[simp]: "normfloat x = x" by transfer rule
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   494
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   495
lemma compute_normfloat[code]: "normfloat (Float m e) =
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   496
  (if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1))
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   497
                           else if m = 0 then 0 else Float m e)"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   498
  by transfer (auto simp add: powr_add zmod_eq_0_iff)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   499
hide_fact (open) compute_normfloat
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   500
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   501
lemma compute_float_numeral[code_abbrev]: "Float (numeral k) 0 = numeral k"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   502
  by transfer simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   503
hide_fact (open) compute_float_numeral
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   504
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   505
lemma compute_float_neg_numeral[code_abbrev]: "Float (- numeral k) 0 = - numeral k"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   506
  by transfer simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   507
hide_fact (open) compute_float_neg_numeral
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   508
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   509
lemma compute_float_uminus[code]: "- Float m1 e1 = Float (- m1) e1"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   510
  by transfer simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   511
hide_fact (open) compute_float_uminus
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   512
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   513
lemma compute_float_times[code]: "Float m1 e1 * Float m2 e2 = Float (m1 * m2) (e1 + e2)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   514
  by transfer (simp add: field_simps powr_add)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   515
hide_fact (open) compute_float_times
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   516
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   517
lemma compute_float_plus[code]: "Float m1 e1 + Float m2 e2 =
54783
25860d89a044 Float: prevent unnecessary large numbers when adding 0
immler
parents: 54782
diff changeset
   518
  (if m1 = 0 then Float m2 e2 else if m2 = 0 then Float m1 e1 else
25860d89a044 Float: prevent unnecessary large numbers when adding 0
immler
parents: 54782
diff changeset
   519
  if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   520
              else Float (m2 + m1 * 2^nat (e1 - e2)) e2)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   521
  by transfer (simp add: field_simps powr_realpow[symmetric] powr_divide2[symmetric])
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   522
hide_fact (open) compute_float_plus
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   523
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   524
lemma compute_float_minus[code]: fixes f g::float shows "f - g = f + (-g)"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   525
  by simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   526
hide_fact (open) compute_float_minus
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   527
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   528
lemma compute_float_sgn[code]: "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   529
  by transfer (simp add: sgn_times)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   530
hide_fact (open) compute_float_sgn
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   531
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
   532
lift_definition is_float_pos :: "float \<Rightarrow> bool" is "op < 0 :: real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   533
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   534
lemma compute_is_float_pos[code]: "is_float_pos (Float m e) \<longleftrightarrow> 0 < m"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   535
  by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0])
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   536
hide_fact (open) compute_is_float_pos
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   537
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   538
lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (b - a)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   539
  by transfer (simp add: field_simps)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   540
hide_fact (open) compute_float_less
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   541
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
   542
lift_definition is_float_nonneg :: "float \<Rightarrow> bool" is "op \<le> 0 :: real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   543
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   544
lemma compute_is_float_nonneg[code]: "is_float_nonneg (Float m e) \<longleftrightarrow> 0 \<le> m"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   545
  by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0])
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   546
hide_fact (open) compute_is_float_nonneg
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   547
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   548
lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (b - a)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   549
  by transfer (simp add: field_simps)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   550
hide_fact (open) compute_float_le
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   551
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
   552
lift_definition is_float_zero :: "float \<Rightarrow> bool"  is "op = 0 :: real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   553
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   554
lemma compute_is_float_zero[code]: "is_float_zero (Float m e) \<longleftrightarrow> 0 = m"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   555
  by transfer (auto simp add: is_float_zero_def)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   556
hide_fact (open) compute_is_float_zero
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   557
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   558
lemma compute_float_abs[code]: "abs (Float m e) = Float (abs m) e"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   559
  by transfer (simp add: abs_mult)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   560
hide_fact (open) compute_float_abs
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   561
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   562
lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   563
  by transfer simp
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   564
hide_fact (open) compute_float_eq
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   565
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   566
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   567
subsection {* Lemmas for types @{typ real}, @{typ nat}, @{typ int}*}
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   568
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   569
lemmas real_of_ints =
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   570
  real_of_int_zero
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   571
  real_of_one
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   572
  real_of_int_add
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   573
  real_of_int_minus
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   574
  real_of_int_diff
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   575
  real_of_int_mult
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   576
  real_of_int_power
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   577
  real_numeral
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   578
lemmas real_of_nats =
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   579
  real_of_nat_zero
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   580
  real_of_nat_one
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   581
  real_of_nat_1
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   582
  real_of_nat_add
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   583
  real_of_nat_mult
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   584
  real_of_nat_power
58989
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
   585
  real_of_nat_numeral
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   586
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   587
lemmas int_of_reals = real_of_ints[symmetric]
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   588
lemmas nat_of_reals = real_of_nats[symmetric]
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   589
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   590
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   591
subsection {* Rounding Real Numbers *}
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   592
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   593
definition round_down :: "int \<Rightarrow> real \<Rightarrow> real" where
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   594
  "round_down prec x = floor (x * 2 powr prec) * 2 powr -prec"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   595
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   596
definition round_up :: "int \<Rightarrow> real \<Rightarrow> real" where
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   597
  "round_up prec x = ceiling (x * 2 powr prec) * 2 powr -prec"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   598
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   599
lemma round_down_float[simp]: "round_down prec x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   600
  unfolding round_down_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   601
  by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   602
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   603
lemma round_up_float[simp]: "round_up prec x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   604
  unfolding round_up_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   605
  by (auto intro!: times_float simp: real_of_int_minus[symmetric] simp del: real_of_int_minus)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   606
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   607
lemma round_up: "x \<le> round_up prec x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   608
  by (simp add: powr_minus_divide le_divide_eq round_up_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   609
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   610
lemma round_down: "round_down prec x \<le> x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   611
  by (simp add: powr_minus_divide divide_le_eq round_down_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   612
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   613
lemma round_up_0[simp]: "round_up p 0 = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   614
  unfolding round_up_def by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   615
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   616
lemma round_down_0[simp]: "round_down p 0 = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   617
  unfolding round_down_def by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   618
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   619
lemma round_up_diff_round_down:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   620
  "round_up prec x - round_down prec x \<le> 2 powr -prec"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   621
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   622
  have "round_up prec x - round_down prec x =
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   623
    (ceiling (x * 2 powr prec) - floor (x * 2 powr prec)) * 2 powr -prec"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   624
    by (simp add: round_up_def round_down_def field_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   625
  also have "\<dots> \<le> 1 * 2 powr -prec"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   626
    by (rule mult_mono)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   627
       (auto simp del: real_of_int_diff
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   628
             simp: real_of_int_diff[symmetric] real_of_int_le_one_cancel_iff ceiling_diff_floor_le_1)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   629
  finally show ?thesis by simp
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   630
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   631
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   632
lemma round_down_shift: "round_down p (x * 2 powr k) = 2 powr k * round_down (p + k) x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   633
  unfolding round_down_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   634
  by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   635
    (simp add: powr_add[symmetric])
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   636
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   637
lemma round_up_shift: "round_up p (x * 2 powr k) = 2 powr k * round_up (p + k) x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   638
  unfolding round_up_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   639
  by (simp add: powr_add powr_mult field_simps powr_divide2[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   640
    (simp add: powr_add[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   641
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   642
lemma round_up_uminus_eq: "round_up p (-x) = - round_down p x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   643
  and round_down_uminus_eq: "round_down p (-x) = - round_up p x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   644
  by (auto simp: round_up_def round_down_def ceiling_def)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   645
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   646
lemma round_up_mono: "x \<le> y \<Longrightarrow> round_up p x \<le> round_up p y"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   647
  by (auto intro!: ceiling_mono simp: round_up_def)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   648
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   649
lemma round_up_le1:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   650
  assumes "x \<le> 1" "prec \<ge> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   651
  shows "round_up prec x \<le> 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   652
proof -
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   653
  have "real \<lceil>x * 2 powr prec\<rceil> \<le> real \<lceil>2 powr real prec\<rceil>"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   654
    using assms by (auto intro!: ceiling_mono)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   655
  also have "\<dots> = 2 powr prec" using assms by (auto simp: powr_int intro!: exI[where x="2^nat prec"])
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   656
  finally show ?thesis
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   657
    by (simp add: round_up_def) (simp add: powr_minus inverse_eq_divide)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   658
qed
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   659
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   660
lemma round_up_less1:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   661
  assumes "x < 1 / 2" "p > 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   662
  shows "round_up p x < 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   663
proof -
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   664
  have "x * 2 powr p < 1 / 2 * 2 powr p"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   665
    using assms by simp
58989
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
   666
  also have "\<dots> \<le> 2 powr p - 1" using `p > 0`
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
   667
    by (auto simp: powr_divide2[symmetric] powr_int field_simps self_le_power)
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
   668
  finally show ?thesis using `p > 0`
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
   669
    by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_eq)
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   670
qed
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   671
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   672
lemma round_down_ge1:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   673
  assumes x: "x \<ge> 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   674
  assumes prec: "p \<ge> - log 2 x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   675
  shows "1 \<le> round_down p x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   676
proof cases
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   677
  assume nonneg: "0 \<le> p"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   678
  have "2 powr p = real \<lfloor>2 powr real p\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   679
    using nonneg by (auto simp: powr_int)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   680
  also have "\<dots> \<le> real \<lfloor>x * 2 powr p\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   681
    using assms by (auto intro!: floor_mono)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   682
  finally show ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   683
    by (simp add: round_down_def) (simp add: powr_minus inverse_eq_divide)
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   684
next
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   685
  assume neg: "\<not> 0 \<le> p"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   686
  have "x = 2 powr (log 2 x)"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   687
    using x by simp
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   688
  also have "2 powr (log 2 x) \<ge> 2 powr - p"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   689
    using prec by auto
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   690
  finally have x_le: "x \<ge> 2 powr -p" .
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   691
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   692
  from neg have "2 powr real p \<le> 2 powr 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   693
    by (intro powr_mono) auto
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   694
  also have "\<dots> \<le> \<lfloor>2 powr 0::real\<rfloor>" by simp
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   695
  also have "... \<le> \<lfloor>x * 2 powr (real p)\<rfloor>" 
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   696
    unfolding real_of_int_le_iff
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   697
    using x x_le by (intro floor_mono) (simp add: powr_minus_divide field_simps)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   698
  finally show ?thesis
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   699
    using prec x
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   700
    by (simp add: round_down_def powr_minus_divide pos_le_divide_eq)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   701
qed
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   702
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   703
lemma round_up_le0: "x \<le> 0 \<Longrightarrow> round_up p x \<le> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   704
  unfolding round_up_def
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   705
  by (auto simp: field_simps mult_le_0_iff zero_le_mult_iff)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   706
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   707
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   708
subsection {* Rounding Floats *}
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   709
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   710
definition div_twopow::"int \<Rightarrow> nat \<Rightarrow> int" where [simp]: "div_twopow x n = x div (2 ^ n)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   711
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   712
definition mod_twopow::"int \<Rightarrow> nat \<Rightarrow> int" where [simp]: "mod_twopow x n = x mod (2 ^ n)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   713
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   714
lemma compute_div_twopow[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   715
  "div_twopow x n = (if x = 0 \<or> x = -1 \<or> n = 0 then x else div_twopow (x div 2) (n - 1))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   716
  by (cases n) (auto simp: zdiv_zmult2_eq div_eq_minus1)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   717
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   718
lemma compute_mod_twopow[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   719
  "mod_twopow x n = (if n = 0 then 0 else x mod 2 + 2 * mod_twopow (x div 2) (n - 1))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   720
  by (cases n) (auto simp: zmod_zmult2_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   721
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   722
lift_definition float_up :: "int \<Rightarrow> float \<Rightarrow> float" is round_up by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   723
declare float_up.rep_eq[simp]
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   724
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   725
lemma round_up_correct:
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   726
  shows "round_up e f - f \<in> {0..2 powr -e}"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   727
unfolding atLeastAtMost_iff
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   728
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   729
  have "round_up e f - f \<le> round_up e f - round_down e f" using round_down by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   730
  also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   731
  finally show "round_up e f - f \<le> 2 powr - (real e)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   732
    by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   733
qed (simp add: algebra_simps round_up)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   734
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   735
lemma float_up_correct:
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   736
  shows "real (float_up e f) - real f \<in> {0..2 powr -e}"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   737
  by transfer (rule round_up_correct)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   738
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   739
lift_definition float_down :: "int \<Rightarrow> float \<Rightarrow> float" is round_down by simp
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   740
declare float_down.rep_eq[simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   741
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   742
lemma round_down_correct:
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   743
  shows "f - (round_down e f) \<in> {0..2 powr -e}"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   744
unfolding atLeastAtMost_iff
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   745
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   746
  have "f - round_down e f \<le> round_up e f - round_down e f" using round_up by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   747
  also have "\<dots> \<le> 2 powr -e" using round_up_diff_round_down by simp
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   748
  finally show "f - round_down e f \<le> 2 powr - (real e)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   749
    by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   750
qed (simp add: algebra_simps round_down)
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   751
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   752
lemma float_down_correct:
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   753
  shows "real f - real (float_down e f) \<in> {0..2 powr -e}"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   754
  by transfer (rule round_down_correct)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   755
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   756
lemma compute_float_down[code]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   757
  "float_down p (Float m e) =
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   758
    (if p + e < 0 then Float (div_twopow m (nat (-(p + e)))) (-p) else Float m e)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   759
proof cases
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   760
  assume "p + e < 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   761
  hence "real ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   762
    using powr_realpow[of 2 "nat (-(p + e))"] by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   763
  also have "... = 1 / 2 powr p / 2 powr e"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   764
    unfolding powr_minus_divide real_of_int_minus by (simp add: powr_add)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   765
  finally show ?thesis
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   766
    using `p + e < 0`
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   767
    by transfer (simp add: ac_simps round_down_def floor_divide_eq_div[symmetric])
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   768
next
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   769
  assume "\<not> p + e < 0"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   770
  then have r: "real e + real p = real (nat (e + p))" by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   771
  have r: "\<lfloor>(m * 2 powr e) * 2 powr real p\<rfloor> = (m * 2 powr e) * 2 powr real p"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   772
    by (auto intro: exI[where x="m*2^nat (e+p)"]
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   773
             simp add: ac_simps powr_add[symmetric] r powr_realpow)
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   774
  with `\<not> p + e < 0` show ?thesis
57862
8f074e6e22fc tuned proofs;
wenzelm
parents: 57512
diff changeset
   775
    by transfer (auto simp add: round_down_def field_simps powr_add powr_minus)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   776
qed
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   777
hide_fact (open) compute_float_down
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   778
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   779
lemma abs_round_down_le: "\<bar>f - (round_down e f)\<bar> \<le> 2 powr -e"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   780
  using round_down_correct[of f e] by simp
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   781
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   782
lemma abs_round_up_le: "\<bar>f - (round_up e f)\<bar> \<le> 2 powr -e"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   783
  using round_up_correct[of e f] by simp
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   784
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   785
lemma round_down_nonneg: "0 \<le> s \<Longrightarrow> 0 \<le> round_down p s"
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56479
diff changeset
   786
  by (auto simp: round_down_def)
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   787
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   788
lemma ceil_divide_floor_conv:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   789
assumes "b \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   790
shows "\<lceil>real a / real b\<rceil> = (if b dvd a then a div b else \<lfloor>real a / real b\<rfloor> + 1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   791
proof cases
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   792
  assume "\<not> b dvd a"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   793
  hence "a mod b \<noteq> 0" by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   794
  hence ne: "real (a mod b) / real b \<noteq> 0" using `b \<noteq> 0` by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   795
  have "\<lceil>real a / real b\<rceil> = \<lfloor>real a / real b\<rfloor> + 1"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   796
  apply (rule ceiling_eq) apply (auto simp: floor_divide_eq_div[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   797
  proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   798
    have "real \<lfloor>real a / real b\<rfloor> \<le> real a / real b" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   799
    moreover have "real \<lfloor>real a / real b\<rfloor> \<noteq> real a / real b"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   800
    apply (subst (2) real_of_int_div_aux) unfolding floor_divide_eq_div using ne `b \<noteq> 0` by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   801
    ultimately show "real \<lfloor>real a / real b\<rfloor> < real a / real b" by arith
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   802
  qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   803
  thus ?thesis using `\<not> b dvd a` by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   804
qed (simp add: ceiling_def real_of_int_minus[symmetric] divide_minus_left[symmetric]
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
   805
  floor_divide_eq_div dvd_neg_div del: divide_minus_left real_of_int_minus)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   806
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   807
lemma compute_float_up[code]:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   808
  "float_up p x = - float_down p (-x)"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   809
  by transfer (simp add: round_down_uminus_eq)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   810
hide_fact (open) compute_float_up
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   811
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   812
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   813
subsection {* Compute bitlen of integers *}
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   814
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   815
definition bitlen :: "int \<Rightarrow> int" where
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   816
  "bitlen a = (if a > 0 then \<lfloor>log 2 a\<rfloor> + 1 else 0)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   817
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   818
lemma bitlen_nonneg: "0 \<le> bitlen x"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   819
proof -
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   820
  {
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   821
    assume "0 > x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   822
    have "-1 = log 2 (inverse 2)" by (subst log_inverse) simp_all
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   823
    also have "... < log 2 (-x)" using `0 > x` by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   824
    finally have "-1 < log 2 (-x)" .
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   825
  } thus "0 \<le> bitlen x" unfolding bitlen_def by (auto intro!: add_nonneg_nonneg)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   826
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   827
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   828
lemma bitlen_bounds:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   829
  assumes "x > 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   830
  shows "2 ^ nat (bitlen x - 1) \<le> x \<and> x < 2 ^ nat (bitlen x)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   831
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   832
  have "(2::real) ^ nat \<lfloor>log 2 (real x)\<rfloor> = 2 powr real (floor (log 2 (real x)))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   833
    using powr_realpow[symmetric, of 2 "nat \<lfloor>log 2 (real x)\<rfloor>"] `x > 0`
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   834
    using real_nat_eq_real[of "floor (log 2 (real x))"]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   835
    by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   836
  also have "... \<le> 2 powr log 2 (real x)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   837
    by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   838
  also have "... = real x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   839
    using `0 < x` by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   840
  finally have "2 ^ nat \<lfloor>log 2 (real x)\<rfloor> \<le> real x" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   841
  thus "2 ^ nat (bitlen x - 1) \<le> x" using `x > 0`
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   842
    by (simp add: bitlen_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   843
next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   844
  have "x \<le> 2 powr (log 2 x)" using `x > 0` by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   845
  also have "... < 2 ^ nat (\<lfloor>log 2 (real x)\<rfloor> + 1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   846
    apply (simp add: powr_realpow[symmetric])
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   847
    using `x > 0` by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   848
  finally show "x < 2 ^ nat (bitlen x)" using `x > 0`
58989
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
   849
    by (simp add: bitlen_def ac_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   850
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   851
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   852
lemma bitlen_pow2[simp]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   853
  assumes "b > 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   854
  shows "bitlen (b * 2 ^ c) = bitlen b + c"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   855
proof -
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
   856
  from assms have "b * 2 ^ c > 0" by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   857
  thus ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   858
    using floor_add[of "log 2 b" c] assms
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   859
    by (auto simp add: log_mult log_nat_power bitlen_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   860
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   861
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   862
lemma bitlen_Float:
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   863
  fixes m e
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   864
  defines "f \<equiv> Float m e"
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   865
  shows "bitlen (\<bar>mantissa f\<bar>) + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)"
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   866
proof (cases "m = 0")
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   867
  case True
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   868
  then show ?thesis by (simp add: f_def bitlen_def Float_def)
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   869
next
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   870
  case False
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   871
  hence "f \<noteq> float_of 0"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   872
    unfolding real_of_float_eq by (simp add: f_def)
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   873
  hence "mantissa f \<noteq> 0"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   874
    by (simp add: mantissa_noteq_0)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   875
  moreover
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   876
  obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i"
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   877
    by (rule f_def[THEN denormalize_shift, OF `f \<noteq> float_of 0`])
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   878
  ultimately show ?thesis by (simp add: abs_mult)
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   879
qed
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   880
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   881
lemma compute_bitlen[code]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   882
  shows "bitlen x = (if x > 0 then bitlen (x div 2) + 1 else 0)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   883
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   884
  { assume "2 \<le> x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   885
    then have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 (x - x mod 2)\<rfloor>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   886
      by (simp add: log_mult zmod_zdiv_equality')
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   887
    also have "\<dots> = \<lfloor>log 2 (real x)\<rfloor>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   888
    proof cases
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   889
      assume "x mod 2 = 0" then show ?thesis by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   890
    next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   891
      def n \<equiv> "\<lfloor>log 2 (real x)\<rfloor>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   892
      then have "0 \<le> n"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   893
        using `2 \<le> x` by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   894
      assume "x mod 2 \<noteq> 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   895
      with `2 \<le> x` have "x mod 2 = 1" "\<not> 2 dvd x" by (auto simp add: dvd_eq_mod_eq_0)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   896
      with `2 \<le> x` have "x \<noteq> 2^nat n" by (cases "nat n") auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   897
      moreover
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   898
      { have "real (2^nat n :: int) = 2 powr (nat n)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   899
          by (simp add: powr_realpow)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   900
        also have "\<dots> \<le> 2 powr (log 2 x)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   901
          using `2 \<le> x` by (simp add: n_def del: powr_log_cancel)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   902
        finally have "2^nat n \<le> x" using `2 \<le> x` by simp }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   903
      ultimately have "2^nat n \<le> x - 1" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   904
      then have "2^nat n \<le> real (x - 1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   905
        unfolding real_of_int_le_iff[symmetric] by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   906
      { have "n = \<lfloor>log 2 (2^nat n)\<rfloor>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   907
          using `0 \<le> n` by (simp add: log_nat_power)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   908
        also have "\<dots> \<le> \<lfloor>log 2 (x - 1)\<rfloor>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   909
          using `2^nat n \<le> real (x - 1)` `0 \<le> n` `2 \<le> x` by (auto intro: floor_mono)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   910
        finally have "n \<le> \<lfloor>log 2 (x - 1)\<rfloor>" . }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   911
      moreover have "\<lfloor>log 2 (x - 1)\<rfloor> \<le> n"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   912
        using `2 \<le> x` by (auto simp add: n_def intro!: floor_mono)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   913
      ultimately show "\<lfloor>log 2 (x - x mod 2)\<rfloor> = \<lfloor>log 2 x\<rfloor>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   914
        unfolding n_def `x mod 2 = 1` by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   915
    qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   916
    finally have "\<lfloor>log 2 (x div 2)\<rfloor> + 1 = \<lfloor>log 2 x\<rfloor>" . }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   917
  moreover
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   918
  { assume "x < 2" "0 < x"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   919
    then have "x = 1" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   920
    then have "\<lfloor>log 2 (real x)\<rfloor> = 0" by simp }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   921
  ultimately show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   922
    unfolding bitlen_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   923
    by (auto simp: pos_imp_zdiv_pos_iff not_le)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   924
qed
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
   925
hide_fact (open) compute_bitlen
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   926
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   927
lemma float_gt1_scale: assumes "1 \<le> Float m e"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   928
  shows "0 \<le> e + (bitlen m - 1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   929
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   930
  have "0 < Float m e" using assms by auto
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   931
  hence "0 < m" using powr_gt_zero[of 2 e]  
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   932
    apply (auto simp: zero_less_mult_iff)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
   933
    using not_le powr_ge_pzero by blast
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   934
  hence "m \<noteq> 0" by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   935
  show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   936
  proof (cases "0 \<le> e")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   937
    case True thus ?thesis using `0 < m`  by (simp add: bitlen_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   938
  next
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   939
    have "(1::int) < 2" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   940
    case False let ?S = "2^(nat (-e))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   941
    have "inverse (2 ^ nat (- e)) = 2 powr e" using assms False powr_realpow[of 2 "nat (-e)"]
57862
8f074e6e22fc tuned proofs;
wenzelm
parents: 57512
diff changeset
   942
      by (auto simp: powr_minus field_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   943
    hence "1 \<le> real m * inverse ?S" using assms False powr_realpow[of 2 "nat (-e)"]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   944
      by (auto simp: powr_minus)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   945
    hence "1 * ?S \<le> real m * inverse ?S * ?S" by (rule mult_right_mono, auto)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57492
diff changeset
   946
    hence "?S \<le> real m" unfolding mult.assoc by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   947
    hence "?S \<le> m" unfolding real_of_int_le_iff[symmetric] by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   948
    from this bitlen_bounds[OF `0 < m`, THEN conjunct2]
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   949
    have "nat (-e) < (nat (bitlen m))" unfolding power_strict_increasing_iff[OF `1 < 2`, symmetric]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   950
      by (rule order_le_less_trans)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   951
    hence "-e < bitlen m" using False by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   952
    thus ?thesis by auto
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   953
  qed
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   954
qed
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   955
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   956
lemma bitlen_div:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   957
  assumes "0 < m"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   958
  shows "1 \<le> real m / 2^nat (bitlen m - 1)" and "real m / 2^nat (bitlen m - 1) < 2"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   959
proof -
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   960
  let ?B = "2^nat(bitlen m - 1)"
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   961
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   962
  have "?B \<le> m" using bitlen_bounds[OF `0 <m`] ..
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   963
  hence "1 * ?B \<le> real m" unfolding real_of_int_le_iff[symmetric] by auto
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   964
  thus "1 \<le> real m / ?B" by auto
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   965
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   966
  have "m \<noteq> 0" using assms by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   967
  have "0 \<le> bitlen m - 1" using `0 < m` by (auto simp: bitlen_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   968
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   969
  have "m < 2^nat(bitlen m)" using bitlen_bounds[OF `0 <m`] ..
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   970
  also have "\<dots> = 2^nat(bitlen m - 1 + 1)" using `0 < m` by (auto simp: bitlen_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   971
  also have "\<dots> = ?B * 2" unfolding nat_add_distrib[OF `0 \<le> bitlen m - 1` zero_le_one] by auto
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   972
  finally have "real m < 2 * ?B" unfolding real_of_int_less_iff[symmetric] by auto
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   973
  hence "real m / ?B < 2 * ?B / ?B" by (rule divide_strict_right_mono, auto)
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   974
  thus "real m / ?B < 2" by auto
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   975
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   976
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   977
subsection {* Truncating Real Numbers*}
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   978
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   979
definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real" where
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   980
  "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   981
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   982
lemma truncate_down: "truncate_down prec x \<le> x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   983
  using round_down by (simp add: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   984
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   985
lemma truncate_down_le: "x \<le> y \<Longrightarrow> truncate_down prec x \<le> y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   986
  by (rule order_trans[OF truncate_down])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   987
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   988
lemma truncate_down_zero[simp]: "truncate_down prec 0 = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   989
  by (simp add: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   990
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   991
lemma truncate_down_float[simp]: "truncate_down p x \<in> float"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   992
  by (auto simp: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   993
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   994
definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real" where
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   995
  "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1) x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   996
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   997
lemma truncate_up: "x \<le> truncate_up prec x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   998
  using round_up by (simp add: truncate_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   999
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1000
lemma truncate_up_le: "x \<le> y \<Longrightarrow> x \<le> truncate_up prec y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1001
  by (rule order_trans[OF _ truncate_up])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1002
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1003
lemma truncate_up_zero[simp]: "truncate_up prec 0 = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1004
  by (simp add: truncate_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1005
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1006
lemma truncate_up_uminus_eq: "truncate_up prec (-x) = - truncate_down prec x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1007
  and truncate_down_uminus_eq: "truncate_down prec (-x) = - truncate_up prec x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1008
  by (auto simp: truncate_up_def round_up_def truncate_down_def round_down_def ceiling_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1009
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1010
lemma truncate_up_float[simp]: "truncate_up p x \<in> float"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1011
  by (auto simp: truncate_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1012
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1013
lemma mult_powr_eq: "0 < b \<Longrightarrow> b \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> x * b powr y = b powr (y + log b x)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1014
  by (simp_all add: powr_add)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1015
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1016
lemma truncate_down_pos:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1017
  assumes "x > 0" "p > 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1018
  shows "truncate_down p x > 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1019
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1020
  have "0 \<le> log 2 x - real \<lfloor>log 2 x\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1021
    by (simp add: algebra_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1022
  from this assms
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1023
  show ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1024
    by (auto simp: truncate_down_def round_down_def mult_powr_eq
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1025
      intro!: ge_one_powr_ge_zero mult_pos_pos)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1026
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1027
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1028
lemma truncate_down_nonneg: "0 \<le> y \<Longrightarrow> 0 \<le> truncate_down prec y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1029
  by (auto simp: truncate_down_def round_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1030
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1031
lemma truncate_down_ge1: "1 \<le> x \<Longrightarrow> 1 \<le> p \<Longrightarrow> 1 \<le> truncate_down p x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1032
  by (auto simp: truncate_down_def algebra_simps intro!: round_down_ge1 add_mono)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1033
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1034
lemma truncate_up_nonpos: "x \<le> 0 \<Longrightarrow> truncate_up prec x \<le> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1035
  by (auto simp: truncate_up_def round_up_def intro!: mult_nonpos_nonneg)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1036
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1037
lemma truncate_up_le1:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1038
  assumes "x \<le> 1" "1 \<le> p" shows "truncate_up p x \<le> 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1039
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1040
  {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1041
    assume "x \<le> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1042
    with truncate_up_nonpos[OF this, of p] have ?thesis by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1043
  } moreover {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1044
    assume "x > 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1045
    hence le: "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<le> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1046
      using assms by (auto simp: log_less_iff)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1047
    from assms have "1 \<le> int p" by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1048
    from add_mono[OF this le]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1049
    have ?thesis using assms
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1050
      by (simp add: truncate_up_def round_up_le1 add_mono)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1051
  } ultimately show ?thesis by arith
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1052
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1053
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1054
subsection {* Truncating Floats*}
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1055
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1056
lift_definition float_round_up :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_up
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1057
  by (simp add: truncate_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1058
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1059
lemma float_round_up: "real x \<le> real (float_round_up prec x)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1060
  using truncate_up by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1061
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1062
lemma float_round_up_zero[simp]: "float_round_up prec 0 = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1063
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1064
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1065
lift_definition float_round_down :: "nat \<Rightarrow> float \<Rightarrow> float" is truncate_down
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1066
  by (simp add: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1067
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1068
lemma float_round_down: "real (float_round_down prec x) \<le> real x"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1069
  using truncate_down by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1070
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1071
lemma float_round_down_zero[simp]: "float_round_down prec 0 = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1072
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1073
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1074
lemmas float_round_up_le = order_trans[OF _ float_round_up]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1075
  and float_round_down_le = order_trans[OF float_round_down]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1076
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1077
lemma minus_float_round_up_eq: "- float_round_up prec x = float_round_down prec (- x)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1078
  and minus_float_round_down_eq: "- float_round_down prec x = float_round_up prec (- x)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1079
  by (transfer, simp add: truncate_down_uminus_eq truncate_up_uminus_eq)+
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1080
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1081
lemma compute_float_round_down[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1082
  "float_round_down prec (Float m e) = (let d = bitlen (abs m) - int prec in
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1083
    if 0 < d then Float (div_twopow m (nat d)) (e + d)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1084
             else Float m e)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1085
  using Float.compute_float_down[of "prec - bitlen \<bar>m\<bar> - e" m e, symmetric]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1086
  by transfer (simp add: field_simps abs_mult log_mult bitlen_def truncate_down_def
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1087
    cong del: if_weak_cong)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1088
hide_fact (open) compute_float_round_down
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1089
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1090
lemma compute_float_round_up[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1091
  "float_round_up prec x = - float_round_down prec (-x)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1092
  by transfer (simp add: truncate_down_uminus_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1093
hide_fact (open) compute_float_round_up
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1094
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1095
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1096
subsection {* Approximation of positive rationals *}
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1097
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1098
lemma div_mult_twopow_eq: fixes a b::nat shows "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1099
  by (cases "b=0") (simp_all add: div_mult2_eq[symmetric] ac_simps)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1100
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1101
lemma real_div_nat_eq_floor_of_divide:
59984
4f1eccec320c conversion between division on nat/int and division in archmedean fields
haftmann
parents: 59554
diff changeset
  1102
  fixes a b :: nat
4f1eccec320c conversion between division on nat/int and division in archmedean fields
haftmann
parents: 59554
diff changeset
  1103
  shows "a div b = real \<lfloor>a / b\<rfloor>"
4f1eccec320c conversion between division on nat/int and division in archmedean fields
haftmann
parents: 59554
diff changeset
  1104
  by (simp add: floor_divide_of_nat_eq [of a b] real_eq_of_nat)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1105
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1106
definition "rat_precision prec x y = int prec - (bitlen x - bitlen y)"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1107
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1108
lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1109
  is "\<lambda>prec (x::nat) (y::nat). round_down (rat_precision prec x y) (x / y)" by simp
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1110
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1111
lemma compute_lapprox_posrat[code]:
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1112
  fixes prec x y
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1113
  shows "lapprox_posrat prec x y =
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1114
   (let
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1115
       l = rat_precision prec x y;
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1116
       d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1117
    in normfloat (Float d (- l)))"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1118
    unfolding div_mult_twopow_eq
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1119
    by transfer
47615
341fd902ef1c transfer now handles Let
hoelzl
parents: 47608
diff changeset
  1120
       (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1121
             del: two_powr_minus_int_float)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  1122
hide_fact (open) compute_lapprox_posrat
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1123
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1124
lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1125
  is "\<lambda>prec (x::nat) (y::nat). round_up (rat_precision prec x y) (x / y)" by simp
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1126
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1127
lemma compute_rapprox_posrat[code]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1128
  fixes prec x y
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1129
  notes divmod_int_mod_div[simp]
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1130
  defines "l \<equiv> rat_precision prec x y"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1131
  shows "rapprox_posrat prec x y = (let
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1132
     l = l ;
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1133
     X = if 0 \<le> l then (x * 2^nat l, y) else (x, y * 2^nat(-l)) ;
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1134
     (d, m) = divmod_int (fst X) (snd X)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1135
   in normfloat (Float (d + (if m = 0 \<or> y = 0 then 0 else 1)) (- l)))"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1136
proof (cases "y = 0")
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1137
  assume "y = 0" thus ?thesis by transfer simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1138
next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1139
  assume "y \<noteq> 0"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1140
  show ?thesis
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1141
  proof (cases "0 \<le> l")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1142
    assume "0 \<le> l"
56777
wenzelm
parents: 56571
diff changeset
  1143
    def x' \<equiv> "x * 2 ^ nat l"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1144
    have "int x * 2 ^ nat l = x'" by (simp add: x'_def int_mult int_power)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1145
    moreover have "real x * 2 powr real l = real x'"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1146
      by (simp add: powr_realpow[symmetric] `0 \<le> l` x'_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1147
    ultimately show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1148
      using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] `0 \<le> l` `y \<noteq> 0`
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1149
        l_def[symmetric, THEN meta_eq_to_obj_eq]
58834
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58410
diff changeset
  1150
      by transfer (auto simp add: floor_divide_eq_div [symmetric] round_up_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1151
   next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1152
    assume "\<not> 0 \<le> l"
56777
wenzelm
parents: 56571
diff changeset
  1153
    def y' \<equiv> "y * 2 ^ nat (- l)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1154
    from `y \<noteq> 0` have "y' \<noteq> 0" by (simp add: y'_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1155
    have "int y * 2 ^ nat (- l) = y'" by (simp add: y'_def int_mult int_power)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1156
    moreover have "real x * real (2::int) powr real l / real y = x / real y'"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1157
      using `\<not> 0 \<le> l`
57862
8f074e6e22fc tuned proofs;
wenzelm
parents: 57512
diff changeset
  1158
      by (simp add: powr_realpow[symmetric] powr_minus y'_def field_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1159
    ultimately show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1160
      using ceil_divide_floor_conv[of y' x] `\<not> 0 \<le> l` `y' \<noteq> 0` `y \<noteq> 0`
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1161
        l_def[symmetric, THEN meta_eq_to_obj_eq]
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1162
      by transfer
58834
773b378d9313 more simp rules concerning dvd and even/odd
haftmann
parents: 58410
diff changeset
  1163
         (auto simp add: round_up_def ceil_divide_floor_conv floor_divide_eq_div [symmetric])
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1164
  qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1165
qed
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  1166
hide_fact (open) compute_rapprox_posrat
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1167
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1168
lemma rat_precision_pos:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1169
  assumes "0 \<le> x" and "0 < y" and "2 * x < y" and "0 < n"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1170
  shows "rat_precision n (int x) (int y) > 0"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1171
proof -
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1172
  { assume "0 < x" hence "log 2 x + 1 = log 2 (2 * x)" by (simp add: log_mult) }
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1173
  hence "bitlen (int x) < bitlen (int y)" using assms
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1174
    by (simp add: bitlen_def del: floor_add_one)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1175
      (auto intro!: floor_mono simp add: floor_add_one[symmetric] simp del: floor_add floor_add_one)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1176
  thus ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1177
    using assms by (auto intro!: pos_add_strict simp add: field_simps rat_precision_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1178
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1179
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1180
lemma rapprox_posrat_less1:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1181
  shows "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 2 * x < y \<Longrightarrow> 0 < n \<Longrightarrow> real (rapprox_posrat n x y) < 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1182
  by transfer (simp add: rat_precision_pos round_up_less1)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1183
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1184
lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1185
  "\<lambda>prec (x::int) (y::int). round_down (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1186
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1187
lemma compute_lapprox_rat[code]:
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1188
  "lapprox_rat prec x y =
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1189
    (if y = 0 then 0
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1190
    else if 0 \<le> x then
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1191
      (if 0 < y then lapprox_posrat prec (nat x) (nat y)
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1192
      else - (rapprox_posrat prec (nat x) (nat (-y))))
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1193
      else (if 0 < y
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1194
        then - (rapprox_posrat prec (nat (-x)) (nat y))
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1195
        else lapprox_posrat prec (nat (-x)) (nat (-y))))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  1196
  by transfer (auto simp: round_up_def round_down_def ceiling_def ac_simps)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  1197
hide_fact (open) compute_lapprox_rat
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1198
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1199
lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1200
  "\<lambda>prec (x::int) (y::int). round_up (rat_precision prec \<bar>x\<bar> \<bar>y\<bar>) (x / y)" by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1201
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1202
lemma "rapprox_rat = rapprox_posrat"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1203
  by transfer auto
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1204
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1205
lemma "lapprox_rat = lapprox_posrat"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1206
  by transfer auto
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1207
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1208
lemma compute_rapprox_rat[code]:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1209
  "rapprox_rat prec x y = - lapprox_rat prec (-x) y"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1210
  by transfer (simp add: round_down_uminus_eq)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  1211
hide_fact (open) compute_rapprox_rat
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1212
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1213
subsection {* Division *}
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1214
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1215
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)"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1216
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1217
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)"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1218
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1219
lift_definition float_divl :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divl
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1220
  by (simp add: real_divl_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1221
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1222
lemma compute_float_divl[code]:
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1223
  "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1224
proof cases
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1225
  let ?f1 = "real m1 * 2 powr real s1" and ?f2 = "real m2 * 2 powr real s2"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1226
  let ?m = "real m1 / real m2" and ?s = "2 powr real (s1 - s2)"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1227
  assume not_0: "m1 \<noteq> 0 \<and> m2 \<noteq> 0"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1228
  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)"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1229
    by (simp add: abs_mult log_mult rat_precision_def bitlen_def)
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1230
  have eq1: "real m1 * 2 powr real s1 / (real m2 * 2 powr real s2) = ?m * ?s"
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1231
    by (simp add: field_simps powr_divide2[symmetric])
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1232
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1233
  show ?thesis
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1234
    using not_0
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1235
    by (transfer fixing: m1 s1 m2 s2 prec) (unfold eq1 eq2 round_down_shift real_divl_def,
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1236
      simp add: field_simps)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1237
qed (transfer, auto simp: real_divl_def)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  1238
hide_fact (open) compute_float_divl
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1239
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1240
lift_definition float_divr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is real_divr
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1241
  by (simp add: real_divr_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1242
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1243
lemma compute_float_divr[code]:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1244
  "float_divr prec x y = - float_divl prec (-x) y"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1245
  by transfer (simp add: real_divr_def real_divl_def round_down_uminus_eq)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1246
hide_fact (open) compute_float_divr
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1247
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1248
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1249
subsection {* Approximate Power *}
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1250
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1251
lemma div2_less_self[termination_simp]: fixes n::nat shows "odd n \<Longrightarrow> n div 2 < n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1252
  by (simp add: odd_pos)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1253
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1254
fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" where
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1255
  "power_down p x 0 = 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1256
| "power_down p x (Suc n) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1257
    (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2) else truncate_down (Suc p) (x * power_down p x n))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1258
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1259
fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" where
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1260
  "power_up p x 0 = 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1261
| "power_up p x (Suc n) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1262
    (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2) else truncate_up p (x * power_up p x n))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1263
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1264
lift_definition power_up_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_up
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1265
  by (induct_tac rule: power_up.induct) simp_all
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1266
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1267
lift_definition power_down_fl :: "nat \<Rightarrow> float \<Rightarrow> nat \<Rightarrow> float" is power_down
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1268
  by (induct_tac rule: power_down.induct) simp_all
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1269
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1270
lemma power_float_transfer[transfer_rule]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1271
  "(rel_fun pcr_float (rel_fun op = pcr_float)) op ^ op ^"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1272
  unfolding power_def
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1273
  by transfer_prover
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1274
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1275
lemma compute_power_up_fl[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1276
  "power_up_fl p x 0 = 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1277
  "power_up_fl p x (Suc n) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1278
    (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2) else float_round_up p (x * power_up_fl p x n))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1279
  and compute_power_down_fl[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1280
  "power_down_fl p x 0 = 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1281
  "power_down_fl p x (Suc n) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1282
    (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2) else float_round_down (Suc p) (x * power_down_fl p x n))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1283
  unfolding atomize_conj
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1284
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1285
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1286
lemma power_down_pos: "0 < x \<Longrightarrow> 0 < power_down p x n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1287
  by (induct p x n rule: power_down.induct)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1288
    (auto simp del: odd_Suc_div_two intro!: truncate_down_pos)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1289
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1290
lemma power_down_nonneg: "0 \<le> x \<Longrightarrow> 0 \<le> power_down p x n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1291
  by (induct p x n rule: power_down.induct)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1292
    (auto simp del: odd_Suc_div_two intro!: truncate_down_nonneg mult_nonneg_nonneg)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1293
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1294
lemma power_down: "0 \<le> x \<Longrightarrow> power_down p x n \<le> x ^ n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1295
proof (induct p x n rule: power_down.induct)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1296
  case (2 p x n)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1297
  {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1298
    assume "odd n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1299
    hence "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1300
      using 2
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1301
      by (auto intro: power_mono power_down_nonneg simp del: odd_Suc_div_two)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1302
    also have "\<dots> = x ^ (Suc n div 2 * 2)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1303
      by (simp add: power_mult[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1304
    also have "Suc n div 2 * 2 = Suc n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1305
      using `odd n` by presburger
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1306
    finally have ?case
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1307
      using `odd n`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1308
      by (auto intro!: truncate_down_le simp del: odd_Suc_div_two)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1309
  } thus ?case
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1310
    by (auto intro!: truncate_down_le mult_left_mono 2 mult_nonneg_nonneg power_down_nonneg)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1311
qed simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1312
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1313
lemma power_up: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up p x n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1314
proof (induct p x n rule: power_up.induct)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1315
  case (2 p x n)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1316
  {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1317
    assume "odd n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1318
    hence "Suc n = Suc n div 2 * 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1319
      using `odd n` even_Suc by presburger
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1320
    hence "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1321
      by (simp add: power_mult[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1322
    also have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1323
      using 2 `odd n`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1324
      by (auto intro: power_mono simp del: odd_Suc_div_two )
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1325
    finally have ?case
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1326
      using `odd n`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1327
      by (auto intro!: truncate_up_le simp del: odd_Suc_div_two )
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1328
  } thus ?case
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1329
    by (auto intro!: truncate_up_le mult_left_mono 2)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1330
qed simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1331
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1332
lemmas power_up_le = order_trans[OF _ power_up]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1333
  and power_up_less = less_le_trans[OF _ power_up]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1334
  and power_down_le = order_trans[OF power_down]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1335
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1336
lemma power_down_fl: "0 \<le> x \<Longrightarrow> power_down_fl p x n \<le> x ^ n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1337
  by transfer (rule power_down)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1338
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1339
lemma power_up_fl: "0 \<le> x \<Longrightarrow> x ^ n \<le> power_up_fl p x n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1340
  by transfer (rule power_up)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1341
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1342
lemma real_power_up_fl: "real (power_up_fl p x n) = power_up p x n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1343
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1344
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1345
lemma real_power_down_fl: "real (power_down_fl p x n) = power_down p x n"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1346
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1347
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1348
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1349
subsection {* Approximate Addition *}
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1350
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1351
definition "plus_down prec x y = truncate_down prec (x + y)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1352
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1353
definition "plus_up prec x y = truncate_up prec (x + y)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1354
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1355
lemma float_plus_down_float[intro, simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> plus_down p x y \<in> float"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1356
  by (simp add: plus_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1357
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1358
lemma float_plus_up_float[intro, simp]: "x \<in> float \<Longrightarrow> y \<in> float \<Longrightarrow> plus_up p x y \<in> float"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1359
  by (simp add: plus_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1360
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1361
lift_definition float_plus_down::"nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is plus_down ..
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1362
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1363
lift_definition float_plus_up::"nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is plus_up ..
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1364
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1365
lemma plus_down: "plus_down prec x y \<le> x + y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1366
  and plus_up: "x + y \<le> plus_up prec x y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1367
  by (auto simp: plus_down_def truncate_down plus_up_def truncate_up)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1368
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1369
lemma float_plus_down: "real (float_plus_down prec x y) \<le> x + y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1370
  and float_plus_up: "x + y \<le> real (float_plus_up prec x y)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1371
  by (transfer, rule plus_down plus_up)+
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1372
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1373
lemmas plus_down_le = order_trans[OF plus_down]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1374
  and plus_up_le = order_trans[OF _ plus_up]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1375
  and float_plus_down_le = order_trans[OF float_plus_down]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1376
  and float_plus_up_le = order_trans[OF _ float_plus_up]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1377
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1378
lemma compute_plus_up[code]: "plus_up p x y = - plus_down p (-x) (-y)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1379
  using truncate_down_uminus_eq[of p "x + y"]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1380
  by (auto simp: plus_down_def plus_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1381
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1382
lemma
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1383
  truncate_down_log2_eqI:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1384
  assumes "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1385
  assumes "\<lfloor>x * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)\<rfloor> = \<lfloor>y * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1386
  shows "truncate_down p x = truncate_down p y"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1387
  using assms by (auto simp: truncate_down_def round_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1388
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1389
lemma bitlen_eq_zero_iff: "bitlen x = 0 \<longleftrightarrow> x \<le> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1390
  by (clarsimp simp add: bitlen_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1391
    (metis Float.compute_bitlen add.commute bitlen_def bitlen_nonneg less_add_same_cancel2 not_less
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1392
      zero_less_one)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1393
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1394
lemma
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1395
  sum_neq_zeroI:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1396
  fixes a k::real
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1397
  shows "abs a \<ge> k \<Longrightarrow> abs b < k \<Longrightarrow> a + b \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1398
    and "abs a > k \<Longrightarrow> abs b \<le> k \<Longrightarrow> a + b \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1399
  by auto
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1400
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1401
lemma
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1402
  abs_real_le_2_powr_bitlen[simp]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1403
  "\<bar>real m2\<bar> < 2 powr real (bitlen \<bar>m2\<bar>)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1404
proof cases
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1405
  assume "m2 \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1406
  hence "\<bar>m2\<bar> < 2 ^ nat (bitlen \<bar>m2\<bar>)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1407
    using bitlen_bounds[of "\<bar>m2\<bar>"]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1408
    by (auto simp: powr_add bitlen_nonneg)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1409
  thus ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1410
    by (simp add: powr_int bitlen_nonneg real_of_int_less_iff[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1411
qed simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1412
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1413
lemma floor_sum_times_2_powr_sgn_eq:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1414
  fixes ai p q::int
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1415
  and a b::real
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1416
  assumes "a * 2 powr p = ai"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1417
  assumes b_le_1: "abs (b * 2 powr (p + 1)) \<le> 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1418
  assumes leqp: "q \<le> p"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1419
  shows "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(2 * ai + sgn b) * 2 powr (q - p - 1)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1420
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1421
  {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1422
    assume "b = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1423
    hence ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1424
      by (simp add: assms(1)[symmetric] powr_add[symmetric] algebra_simps powr_mult_base)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1425
  } moreover {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1426
    assume "b > 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1427
    hence "b * 2 powr p < abs (b * 2 powr (p + 1))" by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1428
    also note b_le_1
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1429
    finally have b_less_1: "b * 2 powr real p < 1" .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1430
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1431
    from b_less_1 `b > 0` have floor_eq: "\<lfloor>b * 2 powr real p\<rfloor> = 0" "\<lfloor>sgn b / 2\<rfloor> = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1432
      by (simp_all add: floor_eq_iff)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1433
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1434
    have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(a + b) * 2 powr p * 2 powr (q - p)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1435
      by (simp add: algebra_simps powr_realpow[symmetric] powr_add[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1436
    also have "\<dots> = \<lfloor>(ai + b * 2 powr p) * 2 powr (q - p)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1437
      by (simp add: assms algebra_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1438
    also have "\<dots> = \<lfloor>(ai + b * 2 powr p) / real ((2::int) ^ nat (p - q))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1439
      using assms
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1440
      by (simp add: algebra_simps powr_realpow[symmetric] divide_powr_uminus powr_add[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1441
    also have "\<dots> = \<lfloor>ai / real ((2::int) ^ nat (p - q))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1442
      by (simp del: real_of_int_power add: floor_divide_real_eq_div floor_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1443
    finally have "\<lfloor>(a + b) * 2 powr real q\<rfloor> = \<lfloor>real ai / real ((2::int) ^ nat (p - q))\<rfloor>" .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1444
    moreover
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1445
    {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1446
      have "\<lfloor>(2 * ai + sgn b) * 2 powr (real (q - p) - 1)\<rfloor> = \<lfloor>(ai + sgn b / 2) * 2 powr (q - p)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1447
        by (subst powr_divide2[symmetric]) (simp add: field_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1448
      also have "\<dots> = \<lfloor>(ai + sgn b / 2) / real ((2::int) ^ nat (p - q))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1449
        using leqp by (simp add: powr_realpow[symmetric] powr_divide2[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1450
      also have "\<dots> = \<lfloor>ai / real ((2::int) ^ nat (p - q))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1451
        by (simp del: real_of_int_power add: floor_divide_real_eq_div floor_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1452
      finally
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1453
      have "\<lfloor>(2 * ai + (sgn b)) * 2 powr (real (q - p) - 1)\<rfloor> =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1454
          \<lfloor>real ai / real ((2::int) ^ nat (p - q))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1455
        .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1456
    } ultimately have ?thesis by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1457
  } moreover {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1458
    assume "\<not> 0 \<le> b"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1459
    hence "0 > b" by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1460
    hence floor_eq: "\<lfloor>b * 2 powr (real p + 1)\<rfloor> = -1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1461
      using b_le_1
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1462
      by (auto simp: floor_eq_iff algebra_simps pos_divide_le_eq[symmetric] abs_if divide_powr_uminus
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1463
        intro!: mult_neg_pos split: split_if_asm)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1464
    have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(2*a + 2*b) * 2 powr p * 2 powr (q - p - 1)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1465
      by (simp add: algebra_simps powr_realpow[symmetric] powr_add[symmetric] powr_mult_base)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1466
    also have "\<dots> = \<lfloor>(2 * (a * 2 powr p) + 2 * b * 2 powr p) * 2 powr (q - p - 1)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1467
      by (simp add: algebra_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1468
    also have "\<dots> = \<lfloor>(2 * ai + b * 2 powr (p + 1)) / 2 powr (1 - q + p)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1469
      using assms by (simp add: algebra_simps powr_mult_base divide_powr_uminus)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1470
    also have "\<dots> = \<lfloor>(2 * ai + b * 2 powr (p + 1)) / real ((2::int) ^ nat (p - q + 1))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1471
      using assms by (simp add: algebra_simps powr_realpow[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1472
    also have "\<dots> = \<lfloor>(2 * ai - 1) / real ((2::int) ^ nat (p - q + 1))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1473
      using `b < 0` assms
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1474
      by (simp add: floor_divide_eq_div floor_eq floor_divide_real_eq_div
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1475
        del: real_of_int_mult real_of_int_power real_of_int_diff)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1476
    also have "\<dots> = \<lfloor>(2 * ai - 1) * 2 powr (q - p - 1)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1477
      using assms by (simp add: algebra_simps divide_powr_uminus powr_realpow[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1478
    finally have ?thesis using `b < 0` by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1479
  } ultimately show ?thesis by arith
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1480
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1481
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1482
lemma
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1483
  log2_abs_int_add_less_half_sgn_eq:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1484
  fixes ai::int and b::real
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1485
  assumes "abs b \<le> 1/2" "ai \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1486
  shows "\<lfloor>log 2 \<bar>real ai + b\<bar>\<rfloor> = \<lfloor>log 2 \<bar>ai + sgn b / 2\<bar>\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1487
proof cases
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1488
  assume "b = 0" thus ?thesis by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1489
next
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1490
  assume "b \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1491
  def k \<equiv> "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1492
  hence "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k" by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1493
  hence k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1494
    by (simp_all add: floor_log_eq_powr_iff `ai \<noteq> 0`)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1495
  have "k \<ge> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1496
    using assms by (auto simp: k_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1497
  def r \<equiv> "\<bar>ai\<bar> - 2 ^ nat k"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1498
  have r: "0 \<le> r" "r < 2 powr k"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1499
    using `k \<ge> 0` k
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1500
    by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1501
  hence "r \<le> (2::int) ^ nat k - 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1502
    using `k \<ge> 0` by (auto simp: powr_int)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1503
  from this[simplified real_of_int_le_iff[symmetric]] `0 \<le> k`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1504
  have r_le: "r \<le> 2 powr k - 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1505
    by (auto simp: algebra_simps powr_int simp del: real_of_int_le_iff)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1506
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1507
  have "\<bar>ai\<bar> = 2 powr k + r"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1508
    using `k \<ge> 0` by (auto simp: k_def r_def powr_realpow[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1509
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1510
  have pos: "\<And>b::real. abs b < 1 \<Longrightarrow> 0 < 2 powr k + (r + b)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1511
    using `0 \<le> k` `ai \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1512
    by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1513
      split: split_if_asm)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1514
  have less: "\<bar>sgn ai * b\<bar> < 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1515
    and less': "\<bar>sgn (sgn ai * b) / 2\<bar> < 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1516
    using `abs b \<le> _` by (auto simp: abs_if sgn_if split: split_if_asm)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1517
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1518
  have floor_eq: "\<And>b::real. abs b \<le> 1 / 2 \<Longrightarrow>
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1519
      \<lfloor>log 2 (1 + (r + b) / 2 powr k)\<rfloor> = (if r = 0 \<and> b < 0 then -1 else 0)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1520
    using `k \<ge> 0` r r_le
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1521
    by (auto simp: floor_log_eq_powr_iff powr_minus_divide field_simps sgn_if)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1522
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1523
  from `real \<bar>ai\<bar> = _` have "\<bar>ai + b\<bar> = 2 powr k + (r + sgn ai * b)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1524
    using `abs b <= _` `0 \<le> k` r
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1525
    by (auto simp add: sgn_if abs_if)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1526
  also have "\<lfloor>log 2 \<dots>\<rfloor> = \<lfloor>log 2 (2 powr k + r + sgn (sgn ai * b) / 2)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1527
  proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1528
    have "2 powr k + (r + (sgn ai) * b) = 2 powr k * (1 + (r + sgn ai * b) / 2 powr k)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1529
      by (simp add: field_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1530
    also have "\<lfloor>log 2 \<dots>\<rfloor> = k + \<lfloor>log 2 (1 + (r + sgn ai * b) / 2 powr k)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1531
      using pos[OF less]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1532
      by (subst log_mult) (simp_all add: log_mult powr_mult field_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1533
    also
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1534
    let ?if = "if r = 0 \<and> sgn ai * b < 0 then -1 else 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1535
    have "\<lfloor>log 2 (1 + (r + sgn ai * b) / 2 powr k)\<rfloor> = ?if"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1536
      using `abs b <= _`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1537
      by (intro floor_eq) (auto simp: abs_mult sgn_if)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1538
    also
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1539
    have "\<dots> = \<lfloor>log 2 (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1540
      by (subst floor_eq) (auto simp: sgn_if)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1541
    also have "k + \<dots> = \<lfloor>log 2 (2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k))\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1542
      unfolding floor_add2[symmetric]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1543
      using pos[OF less'] `abs b \<le> _`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1544
      by (simp add: field_simps add_log_eq_powr)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1545
    also have "2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1546
        2 powr k + r + sgn (sgn ai * b) / 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1547
      by (simp add: sgn_if field_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1548
    finally show ?thesis .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1549
  qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1550
  also have "2 powr k + r + sgn (sgn ai * b) / 2 = \<bar>ai + sgn b / 2\<bar>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1551
    unfolding `real \<bar>ai\<bar> = _`[symmetric] using `ai \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1552
    by (auto simp: abs_if sgn_if algebra_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1553
  finally show ?thesis .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1554
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1555
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1556
lemma compute_far_float_plus_down:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1557
  fixes m1 e1 m2 e2::int and p::nat
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1558
  defines "k1 \<equiv> p - nat (bitlen \<bar>m1\<bar>)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1559
  assumes H: "bitlen \<bar>m2\<bar> \<le> e1 - e2 - k1 - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1560
  shows "float_plus_down p (Float m1 e1) (Float m2 e2) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1561
    float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1562
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1563
  let ?a = "real (Float m1 e1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1564
  let ?b = "real (Float m2 e2)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1565
  let ?sum = "?a + ?b"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1566
  let ?shift = "real e2 - real e1 + real k1 + 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1567
  let ?m1 = "m1 * 2 ^ Suc k1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1568
  let ?m2 = "m2 * 2 powr ?shift"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1569
  let ?m2' = "sgn m2 / 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1570
  let ?e = "e1 - int k1 - 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1571
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1572
  have sum_eq: "?sum = (?m1 + ?m2) * 2 powr ?e"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1573
    by (auto simp: powr_add[symmetric] powr_mult[symmetric] algebra_simps
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1574
      powr_realpow[symmetric] powr_mult_base)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1575
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1576
  have "\<bar>?m2\<bar> * 2 < 2 powr (bitlen \<bar>m2\<bar> + ?shift + 1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1577
    by (auto simp: field_simps powr_add powr_mult_base powr_numeral powr_divide2[symmetric] abs_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1578
  also have "\<dots> \<le> 2 powr 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1579
    using H by (intro powr_mono) auto
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1580
  finally have abs_m2_less_half: "\<bar>?m2\<bar> < 1 / 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1581
    by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1582
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1583
  hence "\<bar>real m2\<bar> < 2 powr -(?shift + 1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1584
    unfolding powr_minus_divide by (auto simp: bitlen_def field_simps powr_mult_base abs_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1585
  also have "\<dots> \<le> 2 powr real (e1 - e2 - 2)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1586
    by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1587
  finally have b_less_quarter: "\<bar>?b\<bar> < 1/4 * 2 powr real e1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1588
    by (simp add: powr_add field_simps powr_divide2[symmetric] powr_numeral abs_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1589
  also have "1/4 < \<bar>real m1\<bar> / 2" using `m1 \<noteq> 0` by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1590
  finally have b_less_half_a: "\<bar>?b\<bar> < 1/2 * \<bar>?a\<bar>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1591
    by (simp add: algebra_simps powr_mult_base abs_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1592
  hence a_half_less_sum: "\<bar>?a\<bar> / 2 < \<bar>?sum\<bar>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1593
    by (auto simp: field_simps abs_if split: split_if_asm)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1594
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1595
  from b_less_half_a have "\<bar>?b\<bar> < \<bar>?a\<bar>" "\<bar>?b\<bar> \<le> \<bar>?a\<bar>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1596
    by simp_all
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1597
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1598
  have "\<bar>real (Float m1 e1)\<bar> \<ge> 1/4 * 2 powr real e1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1599
    using `m1 \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1600
    by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1601
  hence "?sum \<noteq> 0" using b_less_quarter
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1602
    by (rule sum_neq_zeroI)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1603
  hence "?m1 + ?m2 \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1604
    unfolding sum_eq by (simp add: abs_mult zero_less_mult_iff)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1605
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1606
  have "\<bar>real ?m1\<bar> \<ge> 2 ^ Suc k1" "\<bar>?m2'\<bar> < 2 ^ Suc k1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1607
    using `m1 \<noteq> 0` `m2 \<noteq> 0` by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1608
  hence sum'_nz: "?m1 + ?m2' \<noteq> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1609
    by (intro sum_neq_zeroI)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1610
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1611
  have "\<lfloor>log 2 \<bar>real (Float m1 e1) + real (Float m2 e2)\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> + ?e"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1612
    using `?m1 + ?m2 \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1613
    unfolding floor_add[symmetric] sum_eq
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1614
    by (simp add: abs_mult log_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1615
  also have "\<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + sgn (real m2 * 2 powr ?shift) / 2\<bar>\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1616
    using abs_m2_less_half `m1 \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1617
    by (intro log2_abs_int_add_less_half_sgn_eq) (auto simp: abs_mult)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1618
  also have "sgn (real m2 * 2 powr ?shift) = sgn m2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1619
    by (auto simp: sgn_if zero_less_mult_iff less_not_sym)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1620
  also
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1621
  have "\<bar>?m1 + ?m2'\<bar> * 2 powr ?e = \<bar>?m1 * 2 + sgn m2\<bar> * 2 powr (?e - 1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1622
    by (auto simp: field_simps powr_minus[symmetric] powr_divide2[symmetric] powr_mult_base)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1623
  hence "\<lfloor>log 2 \<bar>?m1 + ?m2'\<bar>\<rfloor> + ?e = \<lfloor>log 2 \<bar>real (Float (?m1 * 2 + sgn m2) (?e - 1))\<bar>\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1624
    using `?m1 + ?m2' \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1625
    unfolding floor_add[symmetric]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1626
    by (simp add: log_add_eq_powr abs_mult_pos)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1627
  finally
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1628
  have "\<lfloor>log 2 \<bar>?sum\<bar>\<rfloor> = \<lfloor>log 2 \<bar>real (Float (?m1*2 + sgn m2) (?e - 1))\<bar>\<rfloor>" .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1629
  hence "plus_down p (Float m1 e1) (Float m2 e2) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1630
      truncate_down p (Float (?m1*2 + sgn m2) (?e - 1))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1631
    unfolding plus_down_def
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1632
  proof (rule truncate_down_log2_eqI)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1633
    let ?f = "(int p - \<lfloor>log 2 \<bar>real (Float m1 e1) + real (Float m2 e2)\<bar>\<rfloor> - 1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1634
    let ?ai = "m1 * 2 ^ (Suc k1)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1635
    have "\<lfloor>(?a + ?b) * 2 powr real ?f\<rfloor> = \<lfloor>(real (2 * ?ai) + sgn ?b) * 2 powr real (?f - - ?e - 1)\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1636
    proof (rule floor_sum_times_2_powr_sgn_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1637
      show "?a * 2 powr real (-?e) = real ?ai"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1638
        by (simp add: powr_add powr_realpow[symmetric] powr_divide2[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1639
      show "\<bar>?b * 2 powr real (-?e + 1)\<bar> \<le> 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1640
        using abs_m2_less_half
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1641
        by (simp add: abs_mult powr_add[symmetric] algebra_simps powr_mult_base)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1642
    next
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1643
      have "e1 + \<lfloor>log 2 \<bar>real m1\<bar>\<rfloor> - 1 = \<lfloor>log 2 \<bar>?a\<bar>\<rfloor> - 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1644
        using `m1 \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1645
        by (simp add: floor_add2[symmetric] algebra_simps log_mult abs_mult del: floor_add2)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1646
      also have "\<dots> \<le> \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1647
        using a_half_less_sum `m1 \<noteq> 0` `?sum \<noteq> 0`
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1648
        unfolding floor_subtract[symmetric]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1649
        by (auto simp add: log_minus_eq_powr powr_minus_divide
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1650
          intro!: floor_mono)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1651
      finally
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1652
      have "int p - \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor> \<le> p - (bitlen \<bar>m1\<bar>) - e1 + 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1653
        by (auto simp: algebra_simps bitlen_def `m1 \<noteq> 0`)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1654
      also have "\<dots> \<le> 1 - ?e"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1655
        using bitlen_nonneg[of "\<bar>m1\<bar>"] by (simp add: k1_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1656
      finally show "?f \<le> - ?e" by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1657
    qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1658
    also have "sgn ?b = sgn m2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1659
      using powr_gt_zero[of 2 e2]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1660
      by (auto simp add: sgn_if zero_less_mult_iff simp del: powr_gt_zero)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1661
    also have "\<lfloor>(real (2 * ?m1) + real (sgn m2)) * 2 powr real (?f - - ?e - 1)\<rfloor> =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1662
        \<lfloor>Float (?m1 * 2 + sgn m2) (?e - 1) * 2 powr ?f\<rfloor>"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1663
      by (simp add: powr_add[symmetric] algebra_simps powr_realpow[symmetric])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1664
    finally
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1665
    show "\<lfloor>(?a + ?b) * 2 powr ?f\<rfloor> = \<lfloor>real (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\<rfloor>" .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1666
  qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1667
  thus ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1668
    by transfer (simp add: plus_down_def ac_simps Let_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1669
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1670
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1671
lemma compute_float_plus_down_naive[code]: "float_plus_down p x y = float_round_down p (x + y)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1672
  by transfer (auto simp: plus_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1673
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1674
lemma compute_float_plus_down[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1675
  fixes p::nat and m1 e1 m2 e2::int
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1676
  shows "float_plus_down p (Float m1 e1) (Float m2 e2) =
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1677
    (if m1 = 0 then float_round_down p (Float m2 e2)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1678
    else if m2 = 0 then float_round_down p (Float m1 e1)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1679
    else (if e1 \<ge> e2 then
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1680
      (let
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1681
        k1 = p - nat (bitlen \<bar>m1\<bar>)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1682
      in
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1683
        if bitlen \<bar>m2\<bar> > e1 - e2 - k1 - 2 then float_round_down p ((Float m1 e1) + (Float m2 e2))
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1684
        else float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2)))
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1685
    else float_plus_down p (Float m2 e2) (Float m1 e1)))"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1686
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1687
  {
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1688
    assume H: "bitlen \<bar>m2\<bar> \<le> e1 - e2 - (p - nat (bitlen \<bar>m1\<bar>)) - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1689
    note compute_far_float_plus_down[OF H]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1690
  }
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1691
  thus ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1692
    by transfer (simp add: Let_def plus_down_def ac_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1693
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1694
hide_fact (open) compute_far_float_plus_down
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1695
hide_fact (open) compute_float_plus_down
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1696
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1697
lemma compute_float_plus_up[code]: "float_plus_up p x y = - float_plus_down p (-x) (-y)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1698
  using truncate_down_uminus_eq[of p "x + y"]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1699
  by transfer (simp add: plus_down_def plus_up_def ac_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1700
hide_fact (open) compute_float_plus_up
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1701
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1702
lemma mantissa_zero[simp]: "mantissa 0 = 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1703
by (metis mantissa_0 zero_float.abs_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1704
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1705
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1706
subsection {* Lemmas needed by Approximate *}
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1707
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1708
lemma Float_num[simp]: shows
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1709
   "real (Float 1 0) = 1" and "real (Float 1 1) = 2" and "real (Float 1 2) = 4" and
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58042
diff changeset
  1710
   "real (Float 1 (- 1)) = 1/2" and "real (Float 1 (- 2)) = 1/4" and "real (Float 1 (- 3)) = 1/8" and
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58042
diff changeset
  1711
   "real (Float (- 1) 0) = -1" and "real (Float (number_of n) 0) = number_of n"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1712
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"]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1713
using powr_realpow[of 2 2] powr_realpow[of 2 3]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1714
using powr_minus[of 2 1] powr_minus[of 2 2] powr_minus[of 2 3]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1715
by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1716
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1717
lemma real_of_Float_int[simp]: "real (Float n 0) = real n" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1718
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1719
lemma float_zero[simp]: "real (Float 0 e) = 0" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1720
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1721
lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1722
by arith
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1723
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1724
lemma lapprox_rat:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1725
  shows "real (lapprox_rat prec x y) \<le> real x / real y"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1726
  using round_down by (simp add: lapprox_rat_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1727
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1728
lemma mult_div_le: fixes a b:: int assumes "b > 0" shows "a \<ge> b * (a div b)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1729
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1730
  from zmod_zdiv_equality'[of a b]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1731
  have "a = b * (a div b) + a mod b" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1732
  also have "... \<ge> b * (a div b) + 0" apply (rule add_left_mono) apply (rule pos_mod_sign)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1733
  using assms by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1734
  finally show ?thesis by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1735
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1736
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1737
lemma lapprox_rat_nonneg:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1738
  fixes n x y
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1739
  assumes "0 \<le> x" and "0 \<le> y"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1740
  shows "0 \<le> real (lapprox_rat n x y)"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1741
  using assms by (auto simp: lapprox_rat_def simp: round_down_nonneg)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1742
31098
73dd67adf90a replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
hoelzl
parents: 31021
diff changeset
  1743
lemma rapprox_rat: "real x / real y \<le> real (rapprox_rat prec x y)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1744
  using round_up by (simp add: rapprox_rat_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1745
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1746
lemma rapprox_rat_le1:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1747
  fixes n x y
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1748
  assumes xy: "0 \<le> x" "0 < y" "x \<le> y"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1749
  shows "real (rapprox_rat n x y) \<le> 1"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1750
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1751
  have "bitlen \<bar>x\<bar> \<le> bitlen \<bar>y\<bar>"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1752
    using xy unfolding bitlen_def by (auto intro!: floor_mono)
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1753
  from this assms show ?thesis
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1754
    by transfer (auto intro!: round_up_le1 simp: rat_precision_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1755
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1756
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1757
lemma rapprox_rat_nonneg_nonpos:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1758
  "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1759
  by transfer (simp add: round_up_le0 divide_nonneg_nonpos)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1760
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1761
lemma rapprox_rat_nonpos_nonneg:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1762
  "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real (rapprox_rat n x y) \<le> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1763
  by transfer (simp add: round_up_le0 divide_nonpos_nonneg)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1764
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1765
lemma real_divl: "real_divl prec x y \<le> x / y"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1766
  by (simp add: real_divl_def round_down)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1767
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1768
lemma real_divr: "x / y \<le> real_divr prec x y"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1769
  using round_up by (simp add: real_divr_def)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1770
31098
73dd67adf90a replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
hoelzl
parents: 31021
diff changeset
  1771
lemma float_divl: "real (float_divl prec x y) \<le> real x / real y"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1772
  by transfer (rule real_divl)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1773
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1774
lemma real_divl_lower_bound:
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1775
  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_divl prec x y"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1776
  by (simp add: real_divl_def round_down_nonneg)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1777
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1778
lemma float_divl_lower_bound:
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1779
  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real (float_divl prec x y)"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1780
  by transfer (rule real_divl_lower_bound)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1781
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1782
lemma exponent_1: "exponent 1 = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1783
  using exponent_float[of 1 0] by (simp add: one_float_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1784
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1785
lemma mantissa_1: "mantissa 1 = 1"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1786
  using mantissa_float[of 1 0] by (simp add: one_float_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1787
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1788
lemma bitlen_1: "bitlen 1 = 1"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1789
  by (simp add: bitlen_def)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1790
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1791
lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1792
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1793
  assume "mantissa x = 0" hence z: "0 = real x" using mantissa_exponent by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1794
  show "x = 0" by (simp add: zero_float_def z)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1795
qed (simp add: zero_float_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1796
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1797
lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1798
proof (cases "x = 0", simp)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1799
  assume "x \<noteq> 0" hence "mantissa x \<noteq> 0" using mantissa_eq_zero_iff by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1800
  have "x = mantissa x * 2 powr (exponent x)" by (rule mantissa_exponent)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1801
  also have "mantissa x \<le> \<bar>mantissa x\<bar>" by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1802
  also have "... \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1803
    using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg `mantissa x \<noteq> 0`
58989
99831590def5 tuned proofs
immler
parents: 58987
diff changeset
  1804
    by (auto simp del: real_of_int_abs simp add: powr_int)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1805
  finally show ?thesis by (simp add: powr_add)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1806
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1807
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1808
lemma real_divl_pos_less1_bound:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1809
  assumes "0 < x" "x \<le> 1" "prec \<ge> 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1810
  shows "1 \<le> real_divl prec 1 x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1811
proof -
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1812
  have "log 2 x \<le> real prec + real \<lfloor>log 2 x\<rfloor>" using `prec \<ge> 1` by arith
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1813
  from this assms show ?thesis
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1814
    by (simp add: real_divl_def log_divide round_down_ge1)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1815
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1816
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1817
lemma float_divl_pos_less1_bound:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1818
  "0 < real x \<Longrightarrow> real x \<le> 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow> 1 \<le> real (float_divl prec 1 x)"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1819
  by (transfer, rule real_divl_pos_less1_bound)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1820
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1821
lemma float_divr: "real x / real y \<le> real (float_divr prec x y)"
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1822
  by transfer (rule real_divr)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1823
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1824
lemma real_divr_pos_less1_lower_bound: assumes "0 < x" and "x \<le> 1" shows "1 \<le> real_divr prec 1 x"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1825
proof -
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1826
  have "1 \<le> 1 / x" using `0 < x` and `x <= 1` by auto
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1827
  also have "\<dots> \<le> real_divr prec 1 x" using real_divr[where x=1 and y=x] by auto
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1828
  finally show ?thesis by auto
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1829
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1830
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1831
lemma float_divr_pos_less1_lower_bound: "0 < x \<Longrightarrow> x \<le> 1 \<Longrightarrow> 1 \<le> float_divr prec 1 x"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1832
  by transfer (rule real_divr_pos_less1_lower_bound)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1833
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1834
lemma real_divr_nonpos_pos_upper_bound:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1835
  "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real_divr prec x y \<le> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1836
  by (simp add: real_divr_def round_up_le0 divide_le_0_iff)
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1837
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1838
lemma float_divr_nonpos_pos_upper_bound:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1839
  "real x \<le> 0 \<Longrightarrow> 0 \<le> real y \<Longrightarrow> real (float_divr prec x y) \<le> 0"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1840
  by transfer (rule real_divr_nonpos_pos_upper_bound)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1841
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1842
lemma real_divr_nonneg_neg_upper_bound:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1843
  "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real_divr prec x y \<le> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1844
  by (simp add: real_divr_def round_up_le0 divide_le_0_iff)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1845
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1846
lemma float_divr_nonneg_neg_upper_bound:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1847
  "0 \<le> real x \<Longrightarrow> real y \<le> 0 \<Longrightarrow> real (float_divr prec x y) \<le> 0"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1848
  by transfer (rule real_divr_nonneg_neg_upper_bound)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1849
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1850
lemma truncate_up_nonneg_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1851
  assumes "0 \<le> x" "x \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1852
  shows "truncate_up prec x \<le> truncate_up prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1853
proof -
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1854
  {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1855
    assume "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1856
    hence ?thesis
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1857
      using assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1858
      by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1859
  } moreover {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1860
    assume "0 < x"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1861
    hence "log 2 x \<le> log 2 y" using assms by auto
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1862
    moreover
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1863
    assume "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1864
    ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1865
      unfolding atomize_conj
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1866
      by (metis floor_less_cancel linorder_cases not_le)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1867
    have "truncate_up prec x =
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1868
      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)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1869
      using assms by (simp add: truncate_up_def round_up_def)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1870
    also have "\<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> (2 ^ prec)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1871
    proof (unfold ceiling_le_eq)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1872
      have "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> x * (2 powr real prec / (2 powr log 2 x))"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1873
        using real_of_int_floor_add_one_ge[of "log 2 x"] assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1874
        by (auto simp add: algebra_simps powr_divide2 intro!: mult_left_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1875
      thus "x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> real ((2::int) ^ prec)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1876
        using `0 < x` by (simp add: powr_realpow)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1877
    qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1878
    hence "real \<lceil>x * 2 powr real (int prec - \<lfloor>log 2 x\<rfloor> - 1)\<rceil> \<le> 2 powr int prec"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1879
      by (auto simp: powr_realpow)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1880
    also
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1881
    have "2 powr - real (int prec - \<lfloor>log 2 x\<rfloor> - 1) \<le> 2 powr - real (int prec - \<lfloor>log 2 y\<rfloor>)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1882
      using logless flogless by (auto intro!: floor_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1883
    also have "2 powr real (int prec) \<le> 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>))"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1884
      using assms `0 < x`
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1885
      by (auto simp: algebra_simps)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1886
    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>)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1887
      by simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1888
    also have "\<dots> = 2 powr (log 2 y + real (int prec - \<lfloor>log 2 y\<rfloor>) - real (int prec - \<lfloor>log 2 y\<rfloor>))"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1889
      by (subst powr_add[symmetric]) simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1890
    also have "\<dots> = y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1891
      using `0 < x` assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1892
      by (simp add: powr_add)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1893
    also have "\<dots> \<le> truncate_up prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1894
      by (rule truncate_up)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1895
    finally have ?thesis .
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1896
  } moreover {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1897
    assume "~ 0 < x"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1898
    hence ?thesis
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1899
      using assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1900
      by (auto intro!: truncate_up_le)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1901
  } ultimately show ?thesis
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1902
    by blast
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1903
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1904
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1905
lemma truncate_up_switch_sign_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1906
  assumes "x \<le> 0" "0 \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1907
  shows "truncate_up prec x \<le> truncate_up prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1908
proof -
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1909
  note truncate_up_nonpos[OF `x \<le> 0`]
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1910
  also note truncate_up_le[OF `0 \<le> y`]
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1911
  finally show ?thesis .
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1912
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1913
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1914
lemma truncate_down_zeroprec_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1915
  assumes "0 < x" "x \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1916
  shows "truncate_down 0 x \<le> truncate_down 0 y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1917
proof -
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1918
  have "x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1) = x * inverse (2 powr ((real \<lfloor>log 2 x\<rfloor> + 1)))"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1919
    by (simp add: powr_divide2[symmetric] powr_add powr_minus inverse_eq_divide)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1920
  also have "\<dots> = 2 powr (log 2 x - (real \<lfloor>log 2 x\<rfloor>) - 1)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1921
    using `0 < x`
57862
8f074e6e22fc tuned proofs;
wenzelm
parents: 57512
diff changeset
  1922
    by (auto simp: field_simps powr_add powr_divide2[symmetric])
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1923
  also have "\<dots> < 2 powr 0"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1924
    using real_of_int_floor_add_one_gt
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1925
    unfolding neg_less_iff_less
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1926
    by (intro powr_less_mono) (auto simp: algebra_simps)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1927
  finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> < 1"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1928
    unfolding less_ceiling_eq real_of_int_minus real_of_one
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1929
    by simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1930
  moreover
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1931
  have "0 \<le> \<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor>"
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56479
diff changeset
  1932
    using `x > 0` by auto
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1933
  ultimately have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> \<in> {0 ..< 1}"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1934
    by simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1935
  also have "\<dots> \<subseteq> {0}" by auto
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1936
  finally have "\<lfloor>x * 2 powr (- real \<lfloor>log 2 x\<rfloor> - 1)\<rfloor> = 0" by simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1937
  with assms show ?thesis
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56479
diff changeset
  1938
    by (auto simp: truncate_down_def round_down_def)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1939
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1940
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1941
lemma truncate_down_switch_sign_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1942
  assumes "x \<le> 0" "0 \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1943
  assumes "x \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1944
  shows "truncate_down prec x \<le> truncate_down prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1945
proof -
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1946
  note truncate_down_le[OF `x \<le> 0`]
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1947
  also note truncate_down_nonneg[OF `0 \<le> y`]
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1948
  finally show ?thesis .
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1949
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1950
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1951
lemma truncate_down_nonneg_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1952
  assumes "0 \<le> x" "x \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1953
  shows "truncate_down prec x \<le> truncate_down prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1954
proof -
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1955
  {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1956
    assume "0 < x" "prec = 0"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1957
    with assms have ?thesis
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1958
      by (simp add: truncate_down_zeroprec_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1959
  } moreover {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1960
    assume "~ 0 < x"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1961
    with assms have "x = 0" "0 \<le> y" by simp_all
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1962
    hence ?thesis
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1963
      by (auto intro!: truncate_down_nonneg)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1964
  } moreover {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1965
    assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1966
    hence ?thesis
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1967
      using assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1968
      by (auto simp: truncate_down_def round_down_def intro!: floor_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1969
  } moreover {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1970
    assume "0 < x"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1971
    hence "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y" using assms by auto
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1972
    moreover
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1973
    assume "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1974
    ultimately have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1975
      unfolding atomize_conj abs_of_pos[OF `0 < x`] abs_of_pos[OF `0 < y`]
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1976
      by (metis floor_less_cancel linorder_cases not_le)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1977
    assume "prec \<noteq> 0" hence [simp]: "prec \<ge> Suc 0" by auto
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1978
    have "2 powr (prec - 1) \<le> y * 2 powr real (prec - 1) / (2 powr log 2 y)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1979
      using `0 < y`
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1980
      by simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1981
    also have "\<dots> \<le> y * 2 powr real prec / (2 powr (real \<lfloor>log 2 y\<rfloor> + 1))"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1982
      using `0 \<le> y` `0 \<le> x` assms(2)
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1983
      by (auto intro!: powr_mono divide_left_mono
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1984
        simp: real_of_nat_diff powr_add
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1985
        powr_divide2[symmetric])
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1986
    also have "\<dots> = y * 2 powr real prec / (2 powr real \<lfloor>log 2 y\<rfloor> * 2)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1987
      by (auto simp: powr_add)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1988
    finally have "(2 ^ (prec - 1)) \<le> \<lfloor>y * 2 powr real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1989
      using `0 \<le> y`
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1990
      by (auto simp: powr_divide2[symmetric] le_floor_eq powr_realpow)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1991
    hence "(2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1) \<le> truncate_down prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1992
      by (auto simp: truncate_down_def round_down_def)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1993
    moreover
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1994
    {
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1995
      have "x = 2 powr (log 2 \<bar>x\<bar>)" using `0 < x` by simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1996
      also have "\<dots> \<le> (2 ^ (prec )) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor> - 1)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1997
        using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"]
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1998
        by (auto simp: powr_realpow[symmetric] powr_add[symmetric] algebra_simps)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1999
      also
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2000
      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>)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2001
        using logless flogless `x > 0` `y > 0`
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2002
        by (auto intro!: floor_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2003
      finally have "x \<le> (2 ^ (prec - 1)) * 2 powr - real (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2004
        by (auto simp: powr_realpow[symmetric] powr_divide2[symmetric] assms real_of_nat_diff)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2005
    } ultimately have ?thesis
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2006
      by (metis dual_order.trans truncate_down)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2007
  } ultimately show ?thesis by blast
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2008
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2009
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2010
lemma truncate_down_eq_truncate_up: "truncate_down p x = - truncate_up p (-x)"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2011
  and truncate_up_eq_truncate_down: "truncate_up p x = - truncate_down p (-x)"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2012
  by (auto simp: truncate_up_uminus_eq truncate_down_uminus_eq)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2013
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2014
lemma truncate_down_mono: "x \<le> y \<Longrightarrow> truncate_down p x \<le> truncate_down p y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2015
  apply (cases "0 \<le> x")
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2016
  apply (rule truncate_down_nonneg_mono, assumption+)
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2017
  apply (simp add: truncate_down_eq_truncate_up)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2018
  apply (cases "0 \<le> y")
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2019
  apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2020
  done
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2021
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2022
lemma truncate_up_mono: "x \<le> y \<Longrightarrow> truncate_up p x \<le> truncate_up p y"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2023
  by (simp add: truncate_up_eq_truncate_down truncate_down_mono)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2024
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2025
lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0"
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59984
diff changeset
  2026
 by (auto simp: zero_float_def mult_le_0_iff) (simp add: not_less [symmetric])
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2027
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  2028
lemma real_of_float_pprt[simp]: fixes a::float shows "real (pprt a) = pprt (real a)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2029
  unfolding pprt_def sup_float_def max_def sup_real_def by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2030
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  2031
lemma real_of_float_nprt[simp]: fixes a::float shows "real (nprt a) = nprt (real a)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2032
  unfolding nprt_def inf_float_def min_def inf_real_def by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2033
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
  2034
lift_definition int_floor_fl :: "float \<Rightarrow> int" is floor .
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2035
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2036
lemma compute_int_floor_fl[code]:
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  2037
  "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2038
  by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  2039
hide_fact (open) compute_int_floor_fl
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2040
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2041
lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real (floor x)" by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2042
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2043
lemma compute_floor_fl[code]:
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  2044
  "floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2045
  by transfer (simp add: powr_int int_of_reals floor_divide_eq_div del: real_of_ints)
47621
4cf6011fb884 hide code generation facts in the Float theory, they are only exported for Approximation
hoelzl
parents: 47615
diff changeset
  2046
hide_fact (open) compute_floor_fl
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2047
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2048
lemma floor_fl: "real (floor_fl x) \<le> real x" by transfer simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2049
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2050
lemma int_floor_fl: "real (int_floor_fl x) \<le> real x" by transfer simp
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  2051
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2052
lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0"
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2053
proof (cases "floor_fl x = float_of 0")
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2054
  case True
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2055
  then show ?thesis by (simp add: floor_fl_def)
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2056
next
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2057
  case False
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2058
  have eq: "floor_fl x = Float \<lfloor>real x\<rfloor> 0" by transfer simp
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2059
  obtain i where "\<lfloor>real x\<rfloor> = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i"
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2060
    by (rule denormalize_shift[OF eq[THEN eq_reflection] False])
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2061
  then show ?thesis by simp
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2062
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2063
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2064
lemma compute_mantissa[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2065
  "mantissa (Float m e) = (if m = 0 then 0 else if 2 dvd m then mantissa (normfloat (Float m e)) else m)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2066
  by (auto simp: mantissa_float Float.abs_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2067
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2068
lemma compute_exponent[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2069
  "exponent (Float m e) = (if m = 0 then 0 else if 2 dvd m then exponent (normfloat (Float m e)) else e)"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2070
  by (auto simp: exponent_float Float.abs_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2071
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2072
end
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2073