src/HOL/Library/Float.thy
author haftmann
Sat, 22 Jun 2019 06:25:34 +0000
changeset 70355 a80ad0ac0837
parent 70347 e5cd5471c18a
child 70817 dd675800469d
permissions -rw-r--r--
tuned
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
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
     6
section \<open>Floating-Point Numbers\<close>
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
63663
28d1deca302e Extracted floorlog and bitlen to separate theory Log_Nat
nipkow
parents: 63599
diff changeset
     9
imports Log_Nat 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
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    18
setup_lifting type_definition_float
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
    19
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    20
declare real_of_float [code_unfold]
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    21
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    22
lemmas float_of_inject[simp]
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    23
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    24
declare [[coercion "real_of_float :: float \<Rightarrow> real"]]
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
    25
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    26
lemma real_of_float_eq: "f1 = f2 \<longleftrightarrow> real_of_float f1 = real_of_float f2" for f1 f2 :: float
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    27
  unfolding real_of_float_inject ..
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    28
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    29
declare real_of_float_inverse[simp] float_of_inverse [simp]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    30
declare real_of_float [simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    31
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    32
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
    33
subsection \<open>Real operations preserving the representation as floating point number\<close>
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    34
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    35
lemma floatI: "m * 2 powr e = x \<Longrightarrow> x \<in> float" for m e :: int
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    36
  by (auto simp: float_def)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    37
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    38
lemma zero_float[simp]: "0 \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    39
  by (auto simp: float_def)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    40
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    41
lemma one_float[simp]: "1 \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    42
  by (intro floatI[of 1 0]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    43
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    44
lemma numeral_float[simp]: "numeral i \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    45
  by (intro floatI[of "numeral i" 0]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    46
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    47
lemma neg_numeral_float[simp]: "- numeral i \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    48
  by (intro floatI[of "- numeral i" 0]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    49
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    50
lemma real_of_int_float[simp]: "real_of_int x \<in> float" for x :: int
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    51
  by (intro floatI[of x 0]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    52
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    53
lemma real_of_nat_float[simp]: "real x \<in> float" for x :: nat
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    54
  by (intro floatI[of x 0]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    55
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    56
lemma two_powr_int_float[simp]: "2 powr (real_of_int i) \<in> float" for i :: int
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    57
  by (intro floatI[of 1 i]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    58
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    59
lemma two_powr_nat_float[simp]: "2 powr (real i) \<in> float" for i :: nat
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    60
  by (intro floatI[of 1 i]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    61
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    62
lemma two_powr_minus_int_float[simp]: "2 powr - (real_of_int i) \<in> float" for i :: int
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    63
  by (intro floatI[of 1 "-i"]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    64
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    65
lemma two_powr_minus_nat_float[simp]: "2 powr - (real i) \<in> float" for i :: nat
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    66
  by (intro floatI[of 1 "-i"]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    67
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    68
lemma two_powr_numeral_float[simp]: "2 powr numeral i \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    69
  by (intro floatI[of 1 "numeral i"]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    70
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    71
lemma two_powr_neg_numeral_float[simp]: "2 powr - numeral i \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    72
  by (intro floatI[of 1 "- numeral i"]) simp
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    73
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    74
lemma two_pow_float[simp]: "2 ^ n \<in> float"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
    75
  by (intro floatI[of 1 n]) (simp add: powr_realpow)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
    76
45495
c55a07526dbe cleaned up float theories; removed duplicate definitions and theorems
hoelzl
parents: 44766
diff changeset
    77
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    78
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
    79
  unfolding float_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
    80
proof (safe, simp)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    81
  have *: "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    82
    if "e1 \<le> e2" for e1 m1 e2 m2 :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    83
  proof -
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    84
    from that have "m1 * 2 powr e1 + m2 * 2 powr e2 = (m1 + m2 * 2 ^ nat (e2 - e1)) * 2 powr e1"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
    85
      by (simp add: powr_diff field_simps flip: powr_realpow)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    86
    then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    87
      by blast
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    88
  qed
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    89
  fix e1 m1 e2 m2 :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    90
  consider "e2 \<le> e1" | "e1 \<le> e2" by (rule linorder_le_cases)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    91
  then show "\<exists>(m::int) (e::int). m1 * 2 powr e1 + m2 * 2 powr e2 = m * 2 powr e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    92
  proof cases
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    93
    case 1
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    94
    from *[OF this, of m2 m1] show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    95
      by (simp add: ac_simps)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    96
  next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    97
    case 2
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    98
    then show ?thesis by (rule *)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
    99
  qed
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   100
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   101
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   102
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
   103
  apply (auto simp: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
   104
  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
   105
  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
   106
  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
   107
  apply (rule_tac x="e" in exI)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   108
  apply (simp add: field_simps)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   109
  done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   110
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   111
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
   112
  apply (auto simp: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
   113
  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
   114
  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
   115
  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
   116
  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
   117
  apply (simp add: powr_add)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   118
  done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   119
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   120
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
   121
  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
   122
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   123
lemma abs_float[simp]: "x \<in> float \<Longrightarrow> \<bar>x\<bar> \<in> float"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   124
  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
   125
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   126
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
   127
  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
   128
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   129
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
   130
  apply (auto simp add: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
   131
  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
   132
  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
   133
  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
   134
  apply (rule_tac x="e - d" in exI)
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   135
  apply (simp flip: powr_realpow powr_add add: field_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   136
  done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   137
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   138
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
   139
  apply (auto simp add: float_def)
57492
74bf65a1910a Hypsubst preserves equality hypotheses
Thomas Sewell <thomas.sewell@nicta.com.au>
parents: 56777
diff changeset
   140
  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
   141
  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
   142
  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
   143
  apply (rule_tac x="e - d" in exI)
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   144
  apply (simp flip: powr_realpow powr_add add: field_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   145
  done
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   146
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   147
lemma div_numeral_Bit0_float[simp]:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   148
  assumes "x / numeral n \<in> float"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   149
  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
   150
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   151
  have "(x / numeral n) / 2^1 \<in> float"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   152
    by (intro assms div_power_2_float)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   153
  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
   154
    by (induct n) auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   155
  finally show ?thesis .
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   156
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   157
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   158
lemma div_neg_numeral_Bit0_float[simp]:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   159
  assumes "x / numeral n \<in> float"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   160
  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
   161
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   162
  have "- (x / numeral (Num.Bit0 n)) \<in> float"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   163
    using assms by simp
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   164
  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
   165
    by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   166
  finally show ?thesis .
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   167
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   168
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   169
lemma power_float[simp]:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   170
  assumes "a \<in> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   171
  shows "a ^ b \<in> float"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   172
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   173
  from assms obtain m e :: int where "a = m * 2 powr e"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   174
    by (auto simp: float_def)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   175
  then show ?thesis
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   176
    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
   177
      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
   178
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   179
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   180
lift_definition Float :: "int \<Rightarrow> int \<Rightarrow> float" is "\<lambda>(m::int) (e::int). m * 2 powr e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   181
  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
   182
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
   183
62419
c7d6f4dded19 compute_real_of_float has not been used as code equation
immler
parents: 62390
diff changeset
   184
code_datatype Float
c7d6f4dded19 compute_real_of_float has not been used as code equation
immler
parents: 62390
diff changeset
   185
47780
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   186
lemma compute_real_of_float[code]:
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   187
  "real_of_float (Float m e) = (if e \<ge> 0 then m * 2 ^ nat e else m / 2 ^ (nat (-e)))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   188
  by (simp add: powr_int)
47780
3357688660ff add code equation for real_of_float
hoelzl
parents: 47621
diff changeset
   189
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   190
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   191
subsection \<open>Arithmetic operations on floating point numbers\<close>
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   192
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   193
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
   194
begin
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   195
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   196
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
   197
declare zero_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   198
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   199
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
   200
declare one_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   201
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   202
lift_definition plus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "(+)" 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
   203
declare plus_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   204
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68406
diff changeset
   205
lift_definition times_float :: "float \<Rightarrow> float \<Rightarrow> float" is "(*)" 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
   206
declare times_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   207
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   208
lift_definition minus_float :: "float \<Rightarrow> float \<Rightarrow> float" is "(-)" 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
   209
declare minus_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   210
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   211
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
   212
declare uminus_float.rep_eq[simp]
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   213
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   214
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
   215
declare abs_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   216
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   217
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
   218
declare sgn_float.rep_eq[simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   219
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   220
lift_definition equal_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "(=) :: real \<Rightarrow> real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   221
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   222
lift_definition less_eq_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "(\<le>)" .
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   223
declare less_eq_float.rep_eq[simp]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   224
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   225
lift_definition less_float :: "float \<Rightarrow> float \<Rightarrow> bool" is "(<)" .
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
   226
declare less_float.rep_eq[simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   227
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   228
instance
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   229
  by standard (transfer; fastforce simp add: field_simps intro: mult_left_mono mult_right_mono)+
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   230
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   231
end
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   232
61639
6ef461bee3fa new conversion theorems for int, nat to float
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   233
lemma real_of_float [simp]: "real_of_float (of_nat n) = of_nat n"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   234
  by (induct n) simp_all
61639
6ef461bee3fa new conversion theorems for int, nat to float
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   235
6ef461bee3fa new conversion theorems for int, nat to float
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   236
lemma real_of_float_of_int_eq [simp]: "real_of_float (of_int z) = of_int z"
6ef461bee3fa new conversion theorems for int, nat to float
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   237
  by (cases z rule: int_diff_cases) (simp_all add: of_rat_diff)
6ef461bee3fa new conversion theorems for int, nat to float
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   238
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   239
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
   240
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   241
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   242
lemma real_of_float_power[simp]: "real_of_float (f^n) = real_of_float f^n" for f :: float
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   243
  by (induct n) simp_all
45495
c55a07526dbe cleaned up float theories; removed duplicate definitions and theorems
hoelzl
parents: 44766
diff changeset
   244
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   245
lemma real_of_float_min: "real_of_float (min x y) = min (real_of_float x) (real_of_float y)"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   246
  and real_of_float_max: "real_of_float (max x y) = max (real_of_float x) (real_of_float y)"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   247
  for x y :: float
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   248
  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
   249
53215
5e47c31c6f7c renamed typeclass dense_linorder to unbounded_dense_linorder
hoelzl
parents: 51542
diff changeset
   250
instance float :: unbounded_dense_linorder
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   251
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   252
  fix a b :: float
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   253
  show "\<exists>c. a < c"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   254
    apply (intro exI[of _ "a + 1"])
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   255
    apply transfer
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   256
    apply simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   257
    done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   258
  show "\<exists>c. c < a"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   259
    apply (intro exI[of _ "a - 1"])
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   260
    apply transfer
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   261
    apply simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   262
    done
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   263
  show "\<exists>c. a < c \<and> c < b" if "a < b"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   264
    apply (rule exI[of _ "(a + b) * Float 1 (- 1)"])
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   265
    using that
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   266
    apply transfer
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   267
    apply (simp add: powr_minus)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   268
    done
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   269
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   270
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   271
instantiation float :: lattice_ab_group_add
46573
8c4c5c8dcf7a misc tuning;
wenzelm
parents: 46028
diff changeset
   272
begin
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   273
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   274
definition inf_float :: "float \<Rightarrow> float \<Rightarrow> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   275
  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
   276
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   277
definition sup_float :: "float \<Rightarrow> float \<Rightarrow> float"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   278
  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
   279
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   280
instance
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   281
  by standard (transfer; simp add: inf_float_def sup_float_def real_of_float_min real_of_float_max)+
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   282
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   283
end
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   284
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   285
lemma float_numeral[simp]: "real_of_float (numeral x :: float) = numeral x"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   286
  apply (induct x)
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   287
  apply simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   288
  apply (simp_all only: numeral_Bit0 numeral_Bit1 real_of_float_eq float_of_inverse
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   289
          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
   290
  done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   291
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   292
lemma transfer_numeral [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   293
  "rel_fun (=) pcr_float (numeral :: _ \<Rightarrow> real) (numeral :: _ \<Rightarrow> float)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   294
  by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   295
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   296
lemma float_neg_numeral[simp]: "real_of_float (- numeral x :: float) = - numeral x"
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54230
diff changeset
   297
  by simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   298
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   299
lemma transfer_neg_numeral [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   300
  "rel_fun (=) pcr_float (- numeral :: _ \<Rightarrow> real) (- numeral :: _ \<Rightarrow> float)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   301
  by (simp add: rel_fun_def float.pcr_cr_eq cr_float_def)
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   302
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   303
lemma float_of_numeral: "numeral k = float_of (numeral k)"
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   304
  and float_of_neg_numeral: "- numeral k = float_of (- numeral k)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   305
  unfolding real_of_float_eq by simp_all
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   306
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   307
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   308
subsection \<open>Quickcheck\<close>
58987
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   309
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   310
instantiation float :: exhaustive
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   311
begin
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   312
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   313
definition exhaustive_float where
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   314
  "exhaustive_float f d =
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   315
    Quickcheck_Exhaustive.exhaustive (\<lambda>x. Quickcheck_Exhaustive.exhaustive (\<lambda>y. f (Float x y)) d) d"
58987
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   316
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   317
instance ..
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   318
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   319
end
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   320
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   321
definition (in term_syntax) [code_unfold]:
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   322
  "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
   323
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   324
instantiation float :: full_exhaustive
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   325
begin
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   326
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   327
definition
58987
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   328
  "full_exhaustive_float f d =
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   329
    Quickcheck_Exhaustive.full_exhaustive
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   330
      (\<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
   331
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   332
instance ..
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   333
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   334
end
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   335
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   336
instantiation float :: random
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   337
begin
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   338
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   339
definition "Quickcheck_Random.random i =
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   340
  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
   341
    (\<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
   342
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   343
instance ..
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   344
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   345
end
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   346
119680ebf37c quickcheck setup for float, inspired by rat::{exhaustive,full_exhaustive,random}
immler
parents: 58985
diff changeset
   347
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   348
subsection \<open>Represent floats as unique mantissa and exponent\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46670
diff changeset
   349
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   350
lemma int_induct_abs[case_names less]:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   351
  fixes j :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   352
  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
   353
  shows "P j"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   354
proof (induct "nat \<bar>j\<bar>" arbitrary: j rule: less_induct)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   355
  case less
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   356
  show ?case by (rule H[OF less]) simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   357
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   358
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   359
lemma int_cancel_factors:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   360
  fixes n :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   361
  assumes "1 < r"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   362
  shows "n = 0 \<or> (\<exists>k i. n = k * r ^ i \<and> \<not> r dvd k)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   363
proof (induct n rule: int_induct_abs)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   364
  case (less n)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   365
  have "\<exists>k i. n = k * r ^ Suc i \<and> \<not> r dvd k" if "n \<noteq> 0" "n = m * r" for m
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   366
  proof -
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   367
    from that have "\<bar>m \<bar> < \<bar>n\<bar>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   368
      using \<open>1 < r\<close> by (simp add: abs_mult)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   369
    from less[OF this] that show ?thesis by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   370
  qed
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   371
  then show ?case
59554
4044f53326c9 inlined rules to free user-space from technical names
haftmann
parents: 59487
diff changeset
   372
    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
   373
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   374
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   375
lemma mult_powr_eq_mult_powr_iff_asym:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   376
  fixes m1 m2 e1 e2 :: int
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   377
  assumes m1: "\<not> 2 dvd m1"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   378
    and "e1 \<le> e2"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   379
  shows "m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   380
  (is "?lhs \<longleftrightarrow> ?rhs")
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   381
proof
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   382
  show ?rhs if eq: ?lhs
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   383
  proof -
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   384
    have "m1 \<noteq> 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   385
      using m1 unfolding dvd_def by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   386
    from \<open>e1 \<le> e2\<close> eq have "m1 = m2 * 2 powr nat (e2 - e1)"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
   387
      by (simp add: powr_diff field_simps)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   388
    also have "\<dots> = m2 * 2^nat (e2 - e1)"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   389
      by (simp add: powr_realpow)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   390
    finally have m1_eq: "m1 = m2 * 2^nat (e2 - e1)"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61639
diff changeset
   391
      by linarith
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   392
    with m1 have "m1 = m2"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   393
      by (cases "nat (e2 - e1)") (auto simp add: dvd_def)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   394
    then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   395
      using eq \<open>m1 \<noteq> 0\<close> by (simp add: powr_inj)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   396
  qed
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   397
  show ?lhs if ?rhs
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   398
    using that by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   399
qed
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   400
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   401
lemma mult_powr_eq_mult_powr_iff:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   402
  "\<not> 2 dvd m1 \<Longrightarrow> \<not> 2 dvd m2 \<Longrightarrow> m1 * 2 powr e1 = m2 * 2 powr e2 \<longleftrightarrow> m1 = m2 \<and> e1 = e2"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   403
  for m1 m2 e1 e2 :: int
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   404
  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
   405
  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
   406
  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
   407
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   408
lemma floatE_normed:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   409
  assumes x: "x \<in> float"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   410
  obtains (zero) "x = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   411
   | (powr) m e :: int where "x = m * 2 powr e" "\<not> 2 dvd m" "x \<noteq> 0"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   412
proof -
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   413
  have "\<exists>(m::int) (e::int). x = m * 2 powr e \<and> \<not> (2::int) dvd m" if "x \<noteq> 0"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   414
  proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   415
    from x obtain m e :: int where x: "x = m * 2 powr e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   416
      by (auto simp: float_def)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   417
    with \<open>x \<noteq> 0\<close> int_cancel_factors[of 2 m] obtain k i where "m = k * 2 ^ i" "\<not> 2 dvd k"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   418
      by auto
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   419
    with \<open>\<not> 2 dvd k\<close> x show ?thesis
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   420
      apply (rule_tac exI[of _ "k"])
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   421
      apply (rule_tac exI[of _ "e + int i"])
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   422
      apply (simp add: powr_add powr_realpow)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   423
      done
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   424
  qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   425
  with that show thesis by blast
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   426
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   427
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   428
lemma float_normed_cases:
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   429
  fixes f :: float
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   430
  obtains (zero) "f = 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   431
   | (powr) m e :: int where "real_of_float f = m * 2 powr e" "\<not> 2 dvd m" "f \<noteq> 0"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   432
proof (atomize_elim, induct f)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   433
  case (float_of y)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   434
  then show ?case
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   435
    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
   436
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   437
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   438
definition mantissa :: "float \<Rightarrow> int"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   439
  where "mantissa f =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   440
    fst (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) \<or>
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   441
      (f \<noteq> 0 \<and> real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \<and> \<not> 2 dvd fst p))"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   442
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   443
definition exponent :: "float \<Rightarrow> int"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   444
  where "exponent f =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   445
    snd (SOME p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) \<or>
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   446
      (f \<noteq> 0 \<and> real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \<and> \<not> 2 dvd fst p))"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   447
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   448
lemma exponent_0[simp]: "exponent 0 = 0" (is ?E)
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   449
  and mantissa_0[simp]: "mantissa 0 = 0" (is ?M)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   450
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   451
  have "\<And>p::int \<times> int. fst p = 0 \<and> snd p = 0 \<longleftrightarrow> p = (0, 0)"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   452
    by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   453
  then show ?E ?M
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   454
    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
   455
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   456
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   457
lemma mantissa_exponent: "real_of_float f = mantissa f * 2 powr exponent f" (is ?E)
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   458
  and mantissa_not_dvd: "f \<noteq> 0 \<Longrightarrow> \<not> 2 dvd mantissa f" (is "_ \<Longrightarrow> ?D")
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   459
proof cases
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   460
  assume [simp]: "f \<noteq> 0"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   461
  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
   462
  proof (cases f rule: float_normed_cases)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   463
    case zero
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   464
    then show ?thesis by simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   465
  next
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   466
    case (powr m e)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   467
    then have "\<exists>p::int \<times> int. (f = 0 \<and> fst p = 0 \<and> snd p = 0) \<or>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   468
      (f \<noteq> 0 \<and> real_of_float f = real_of_int (fst p) * 2 powr real_of_int (snd p) \<and> \<not> 2 dvd fst p)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   469
      by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   470
    then show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   471
      unfolding exponent_def mantissa_def
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   472
      by (rule someI2_ex) simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   473
  qed
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   474
  then show ?E ?D by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   475
qed simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   476
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   477
lemma mantissa_noteq_0: "f \<noteq> 0 \<Longrightarrow> mantissa f \<noteq> 0"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   478
  using mantissa_not_dvd[of f] by auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   479
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   480
lemma mantissa_eq_zero_iff: "mantissa x = 0 \<longleftrightarrow> x = 0"
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   481
  (is "?lhs \<longleftrightarrow> ?rhs")
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   482
proof
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   483
  show ?rhs if ?lhs
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   484
  proof -
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   485
    from that have z: "0 = real_of_float x"
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   486
      using mantissa_exponent by simp
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   487
    show ?thesis
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   488
      by (simp add: zero_float_def z)
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   489
  qed
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   490
  show ?lhs if ?rhs
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   491
    using that by simp
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   492
qed
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   493
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   494
lemma mantissa_pos_iff: "0 < mantissa x \<longleftrightarrow> 0 < x"
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   495
  by (auto simp: mantissa_exponent sign_simps)
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   496
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   497
lemma mantissa_nonneg_iff: "0 \<le> mantissa x \<longleftrightarrow> 0 \<le> x"
70347
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
   498
  by (auto simp: mantissa_exponent sign_simps)
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   499
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   500
lemma mantissa_neg_iff: "0 > mantissa x \<longleftrightarrow> 0 > x"
70347
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
   501
  by (auto simp: mantissa_exponent sign_simps)
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   502
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
   503
lemma
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   504
  fixes m e :: int
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   505
  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
   506
  assumes dvd: "\<not> 2 dvd m"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   507
  shows mantissa_float: "mantissa f = m" (is "?M")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   508
    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
   509
proof cases
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   510
  assume "m = 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   511
  with dvd show "mantissa f = m" by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   512
next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   513
  assume "m \<noteq> 0"
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   514
  then have f_not_0: "f \<noteq> 0" by (simp add: f_def zero_float_def)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   515
  from mantissa_exponent[of f] have "m * 2 powr e = mantissa f * 2 powr exponent f"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   516
    by (auto simp add: f_def)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   517
  then show ?M ?E
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   518
    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
   519
    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
   520
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   521
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   522
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   523
subsection \<open>Compute arithmetic operations\<close>
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   524
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   525
lemma Float_mantissa_exponent: "Float (mantissa f) (exponent f) = f"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   526
  unfolding real_of_float_eq mantissa_exponent[of f] by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   527
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   528
lemma Float_cases [cases type: float]:
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   529
  fixes f :: float
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   530
  obtains (Float) m e :: int where "f = Float m e"
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   531
  using Float_mantissa_exponent[symmetric]
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   532
  by (atomize_elim) auto
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   533
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   534
lemma denormalize_shift:
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   535
  assumes f_def: "f = Float m e"
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   536
    and not_0: "f \<noteq> 0"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   537
  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
   538
proof
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   539
  from mantissa_exponent[of f] f_def
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   540
  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
   541
    by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   542
  then have eq: "m = mantissa f * 2 powr (exponent f - e)"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
   543
    by (simp add: powr_diff field_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   544
  moreover
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   545
  have "e \<le> exponent f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   546
  proof (rule ccontr)
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   547
    assume "\<not> e \<le> exponent f"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   548
    then have pos: "exponent f < e" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   549
    then have "2 powr (exponent f - e) = 2 powr - real_of_int (e - exponent f)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   550
      by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   551
    also have "\<dots> = 1 / 2^nat (e - exponent f)"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   552
      using pos by (simp flip: powr_realpow add: powr_diff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   553
    finally have "m * 2^nat (e - exponent f) = real_of_int (mantissa f)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   554
      using eq by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   555
    then have "mantissa f = m * 2^nat (e - exponent f)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   556
      by linarith
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   557
    with \<open>exponent f < e\<close> have "2 dvd mantissa f"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   558
      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
   559
      apply (cases "nat (e - exponent f)")
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   560
      apply auto
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   561
      done
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   562
    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
   563
  qed
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   564
  ultimately have "real_of_int m = mantissa f * 2^nat (exponent f - e)"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   565
    by (simp flip: powr_realpow)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   566
  with \<open>e \<le> exponent f\<close>
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   567
  show "m = mantissa f * 2 ^ nat (exponent f - e)"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61639
diff changeset
   568
    by linarith
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61639
diff changeset
   569
  show "e = exponent f - nat (exponent f - e)"
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61762
diff changeset
   570
    using \<open>e \<le> exponent f\<close> by auto
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   571
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   572
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   573
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   574
begin
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   575
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   576
qualified 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
   577
  by transfer simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   578
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   579
qualified lemma compute_float_one[code_unfold, code]: "1 = Float 1 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   580
  by transfer simp
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   581
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   582
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
   583
lemma normloat_id[simp]: "normfloat x = x" by transfer rule
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   584
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   585
qualified lemma compute_normfloat[code]:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   586
  "normfloat (Float m e) =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   587
    (if m mod 2 = 0 \<and> m \<noteq> 0 then normfloat (Float (m div 2) (e + 1))
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   588
     else if m = 0 then 0 else Float m e)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   589
  by transfer (auto simp add: powr_add zmod_eq_0_iff)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   590
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   591
qualified 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
   592
  by transfer simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   593
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   594
qualified 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
   595
  by transfer simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   596
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   597
qualified 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
   598
  by transfer simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   599
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   600
qualified 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
   601
  by transfer (simp add: field_simps powr_add)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   602
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   603
qualified lemma compute_float_plus[code]:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   604
  "Float m1 e1 + Float m2 e2 =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   605
    (if m1 = 0 then Float m2 e2
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   606
     else if m2 = 0 then Float m1 e1
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   607
     else if e1 \<le> e2 then Float (m1 + m2 * 2^nat (e2 - e1)) e1
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   608
     else Float (m2 + m1 * 2^nat (e1 - e2)) e2)"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
   609
  by transfer (simp add: field_simps powr_realpow[symmetric] powr_diff)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   610
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   611
qualified lemma compute_float_minus[code]: "f - g = f + (-g)" for f g :: float
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   612
  by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   613
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   614
qualified lemma compute_float_sgn[code]:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   615
  "sgn (Float m1 e1) = (if 0 < m1 then 1 else if m1 < 0 then -1 else 0)"
64240
eabf80376aab more standardized names
haftmann
parents: 63664
diff changeset
   616
  by transfer (simp add: sgn_mult)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   617
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   618
lift_definition is_float_pos :: "float \<Rightarrow> bool" is "(<) 0 :: real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   619
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   620
qualified 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
   621
  by transfer (auto simp add: zero_less_mult_iff not_le[symmetric, of _ 0])
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   622
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   623
lift_definition is_float_nonneg :: "float \<Rightarrow> bool" is "(\<le>) 0 :: real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   624
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   625
qualified 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
   626
  by transfer (auto simp add: zero_le_mult_iff not_less[symmetric, of _ 0])
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   627
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
   628
lift_definition is_float_zero :: "float \<Rightarrow> bool"  is "(=) 0 :: real \<Rightarrow> bool" .
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   629
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   630
qualified 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
   631
  by transfer (auto simp add: is_float_zero_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   632
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   633
qualified lemma compute_float_abs[code]: "\<bar>Float m e\<bar> = Float \<bar>m\<bar> e"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   634
  by transfer (simp add: abs_mult)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   635
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   636
qualified lemma compute_float_eq[code]: "equal_class.equal f g = is_float_zero (f - g)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   637
  by transfer simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   638
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   639
end
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   640
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   641
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69064
diff changeset
   642
subsection \<open>Lemmas for types \<^typ>\<open>real\<close>, \<^typ>\<open>nat\<close>, \<^typ>\<open>int\<close>\<close>
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   643
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   644
lemmas real_of_ints =
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   645
  of_int_add
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   646
  of_int_minus
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   647
  of_int_diff
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   648
  of_int_mult
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   649
  of_int_power
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   650
  of_int_numeral of_int_neg_numeral
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   651
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   652
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
   653
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   654
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   655
subsection \<open>Rounding Real Numbers\<close>
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   656
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   657
definition round_down :: "int \<Rightarrow> real \<Rightarrow> real"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
   658
  where "round_down prec x = \<lfloor>x * 2 powr prec\<rfloor> * 2 powr -prec"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   659
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   660
definition round_up :: "int \<Rightarrow> real \<Rightarrow> real"
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
   661
  where "round_up prec x = \<lceil>x * 2 powr prec\<rceil> * 2 powr -prec"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   662
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   663
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
   664
  unfolding round_down_def
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   665
  by (auto intro!: times_float simp flip: of_int_minus)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   666
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   667
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
   668
  unfolding round_up_def
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   669
  by (auto intro!: times_float simp flip: of_int_minus)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   670
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   671
lemma round_up: "x \<le> round_up prec x"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   672
  by (simp add: powr_minus_divide le_divide_eq round_up_def ceiling_correct)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   673
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   674
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
   675
  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
   676
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   677
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
   678
  unfolding round_up_def by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   679
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   680
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
   681
  unfolding round_down_def by simp
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   682
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   683
lemma round_up_diff_round_down: "round_up prec x - round_down prec x \<le> 2 powr -prec"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   684
proof -
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   685
  have "round_up prec x - round_down prec x = (\<lceil>x * 2 powr prec\<rceil> - \<lfloor>x * 2 powr prec\<rfloor>) * 2 powr -prec"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   686
    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
   687
  also have "\<dots> \<le> 1 * 2 powr -prec"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   688
    by (rule mult_mono)
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   689
      (auto simp flip: of_int_diff simp: ceiling_diff_floor_le_1)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   690
  finally show ?thesis by simp
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   691
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   692
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   693
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
   694
  unfolding round_down_def
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
   695
  by (simp add: powr_add powr_mult field_simps powr_diff)
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   696
    (simp flip: powr_add)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   697
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   698
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
   699
  unfolding round_up_def
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
   700
  by (simp add: powr_add powr_mult field_simps powr_diff)
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   701
    (simp flip: powr_add)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   702
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   703
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
   704
  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
   705
  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
   706
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   707
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
   708
  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
   709
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   710
lemma round_up_le1:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   711
  assumes "x \<le> 1" "prec \<ge> 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   712
  shows "round_up prec x \<le> 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   713
proof -
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   714
  have "real_of_int \<lceil>x * 2 powr prec\<rceil> \<le> real_of_int \<lceil>2 powr real_of_int prec\<rceil>"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   715
    using assms by (auto intro!: ceiling_mono)
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   716
  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
   717
  finally show ?thesis
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   718
    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
   719
qed
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   720
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   721
lemma round_up_less1:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   722
  assumes "x < 1 / 2" "p > 0"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   723
  shows "round_up p x < 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   724
proof -
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   725
  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
   726
    using assms by simp
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   727
  also have "\<dots> \<le> 2 powr p - 1" using \<open>p > 0\<close>
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
   728
    by (auto simp: powr_diff powr_int field_simps self_le_power)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   729
  finally show ?thesis using \<open>p > 0\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   730
    by (simp add: round_up_def field_simps powr_minus powr_int ceiling_less_iff)
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   731
qed
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   732
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   733
lemma round_down_ge1:
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   734
  assumes x: "x \<ge> 1"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   735
  assumes prec: "p \<ge> - log 2 x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   736
  shows "1 \<le> round_down p x"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   737
proof cases
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   738
  assume nonneg: "0 \<le> p"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   739
  have "2 powr p = real_of_int \<lfloor>2 powr real_of_int p\<rfloor>"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   740
    using nonneg by (auto simp: powr_int)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   741
  also have "\<dots> \<le> real_of_int \<lfloor>x * 2 powr p\<rfloor>"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   742
    using assms by (auto intro!: floor_mono)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   743
  finally show ?thesis
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   744
    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
   745
next
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   746
  assume neg: "\<not> 0 \<le> p"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   747
  have "x = 2 powr (log 2 x)"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   748
    using x by simp
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   749
  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
   750
    using prec by auto
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   751
  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
   752
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   753
  from neg have "2 powr real_of_int p \<le> 2 powr 0"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   754
    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
   755
  also have "\<dots> \<le> \<lfloor>2 powr 0::real\<rfloor>" by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   756
  also have "\<dots> \<le> \<lfloor>x * 2 powr (real_of_int p)\<rfloor>"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   757
    unfolding of_int_le_iff
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   758
    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
   759
  finally show ?thesis
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   760
    using prec x
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   761
    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
   762
qed
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   763
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   764
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
   765
  unfolding round_up_def
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   766
  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
   767
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   768
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   769
subsection \<open>Rounding Floats\<close>
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   770
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   771
definition div_twopow :: "int \<Rightarrow> nat \<Rightarrow> int"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   772
  where [simp]: "div_twopow x n = x div (2 ^ n)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   773
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   774
definition mod_twopow :: "int \<Rightarrow> nat \<Rightarrow> int"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   775
  where [simp]: "mod_twopow x n = x mod (2 ^ n)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   776
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   777
lemma compute_div_twopow[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   778
  "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
   779
  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
   780
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   781
lemma compute_mod_twopow[code]:
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   782
  "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
   783
  by (cases n) (auto simp: zmod_zmult2_eq)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   784
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   785
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
   786
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
   787
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   788
lemma round_up_correct: "round_up e f - f \<in> {0..2 powr -e}"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   789
  unfolding atLeastAtMost_iff
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   790
proof
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   791
  have "round_up e f - f \<le> round_up e f - round_down e f"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   792
    using round_down by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   793
  also have "\<dots> \<le> 2 powr -e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   794
    using round_up_diff_round_down by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   795
  finally show "round_up e f - f \<le> 2 powr - (real_of_int e)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   796
    by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   797
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
   798
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   799
lemma float_up_correct: "real_of_float (float_up e f) - real_of_float f \<in> {0..2 powr -e}"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   800
  by transfer (rule round_up_correct)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   801
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   802
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
   803
declare float_down.rep_eq[simp]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   804
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   805
lemma round_down_correct: "f - (round_down e f) \<in> {0..2 powr -e}"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   806
  unfolding atLeastAtMost_iff
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   807
proof
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   808
  have "f - round_down e f \<le> round_up e f - round_down e f"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   809
    using round_up by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   810
  also have "\<dots> \<le> 2 powr -e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   811
    using round_up_diff_round_down by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   812
  finally show "f - round_down e f \<le> 2 powr - (real_of_int e)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   813
    by simp
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   814
qed (simp add: algebra_simps round_down)
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   815
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   816
lemma float_down_correct: "real_of_float f - real_of_float (float_down e f) \<in> {0..2 powr -e}"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   817
  by transfer (rule round_down_correct)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   818
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   819
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   820
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   821
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   822
qualified lemma compute_float_down[code]:
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   823
  "float_down p (Float m e) =
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   824
    (if p + e < 0 then Float (div_twopow m (nat (-(p + e)))) (-p) else Float m e)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   825
proof (cases "p + e < 0")
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   826
  case True
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   827
  then have "real_of_int ((2::int) ^ nat (-(p + e))) = 2 powr (-(p + e))"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   828
    using powr_realpow[of 2 "nat (-(p + e))"] by simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   829
  also have "\<dots> = 1 / 2 powr p / 2 powr e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   830
    unfolding powr_minus_divide 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
   831
  finally show ?thesis
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   832
    using \<open>p + e < 0\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   833
    apply transfer
70355
haftmann
parents: 70347
diff changeset
   834
    apply (simp add: round_down_def field_simps flip: floor_divide_of_int_eq)
haftmann
parents: 70347
diff changeset
   835
    apply (metis (no_types, hide_lams) Float.rep_eq
haftmann
parents: 70347
diff changeset
   836
      add.inverse_inverse compute_real_of_float diff_minus_eq_add
haftmann
parents: 70347
diff changeset
   837
      floor_divide_of_int_eq int_of_reals(1) linorder_not_le
haftmann
parents: 70347
diff changeset
   838
      minus_add_distrib of_int_eq_numeral_power_cancel_iff powr_add)
haftmann
parents: 70347
diff changeset
   839
    done
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   840
next
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   841
  case False
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   842
  then have r: "real_of_int e + real_of_int p = real (nat (e + p))"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   843
    by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   844
  have r: "\<lfloor>(m * 2 powr e) * 2 powr real_of_int p\<rfloor> = (m * 2 powr e) * 2 powr real_of_int p"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
   845
    by (auto intro: exI[where x="m*2^nat (e+p)"]
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   846
        simp add: ac_simps powr_add[symmetric] r powr_realpow)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   847
  with \<open>\<not> p + e < 0\<close> show ?thesis
57862
8f074e6e22fc tuned proofs;
wenzelm
parents: 57512
diff changeset
   848
    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
   849
qed
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   850
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   851
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
   852
  using round_down_correct[of f e] by simp
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   853
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   854
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
   855
  using round_up_correct[of e f] by simp
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   856
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   857
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
   858
  by (auto simp: round_down_def)
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
   859
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   860
lemma ceil_divide_floor_conv:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   861
  assumes "b \<noteq> 0"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   862
  shows "\<lceil>real_of_int a / real_of_int b\<rceil> =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   863
    (if b dvd a then a div b else \<lfloor>real_of_int a / real_of_int b\<rfloor> + 1)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   864
proof (cases "b dvd a")
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   865
  case True
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   866
  then show ?thesis
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   867
    by (simp add: ceiling_def floor_divide_of_int_eq dvd_neg_div
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   868
             flip: of_int_minus divide_minus_left)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   869
next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   870
  case False
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   871
  then have "a mod b \<noteq> 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   872
    by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   873
  then have ne: "real_of_int (a mod b) / real_of_int b \<noteq> 0"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   874
    using \<open>b \<noteq> 0\<close> by auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   875
  have "\<lceil>real_of_int a / real_of_int b\<rceil> = \<lfloor>real_of_int a / real_of_int b\<rfloor> + 1"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   876
    apply (rule ceiling_eq)
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
   877
    apply (auto simp flip: floor_divide_of_int_eq)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   878
  proof -
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   879
    have "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> \<le> real_of_int a / real_of_int b"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   880
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   881
    moreover have "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> \<noteq> real_of_int a / real_of_int b"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   882
      apply (subst (2) real_of_int_div_aux)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   883
      unfolding floor_divide_of_int_eq
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   884
      using ne \<open>b \<noteq> 0\<close> apply auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   885
      done
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   886
    ultimately show "real_of_int \<lfloor>real_of_int a / real_of_int b\<rfloor> < real_of_int a / real_of_int b" by arith
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   887
  qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   888
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   889
    using \<open>\<not> b dvd a\<close> by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   890
qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   891
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   892
qualified lemma compute_float_up[code]: "float_up p x = - float_down p (-x)"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
   893
  by transfer (simp add: round_down_uminus_eq)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   894
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   895
end
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   896
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   897
63664
nipkow
parents: 63663
diff changeset
   898
lemma bitlen_Float:
nipkow
parents: 63663
diff changeset
   899
  fixes m e
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   900
  defines [THEN meta_eq_to_obj_eq]: "f \<equiv> Float m e"
63664
nipkow
parents: 63663
diff changeset
   901
  shows "bitlen \<bar>mantissa f\<bar> + exponent f = (if m = 0 then 0 else bitlen \<bar>m\<bar> + e)"
nipkow
parents: 63663
diff changeset
   902
proof (cases "m = 0")
nipkow
parents: 63663
diff changeset
   903
  case True
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   904
  then show ?thesis by (simp add: f_def bitlen_alt_def)
63664
nipkow
parents: 63663
diff changeset
   905
next
nipkow
parents: 63663
diff changeset
   906
  case False
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   907
  then have "f \<noteq> 0"
63664
nipkow
parents: 63663
diff changeset
   908
    unfolding real_of_float_eq by (simp add: f_def)
nipkow
parents: 63663
diff changeset
   909
  then have "mantissa f \<noteq> 0"
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   910
    by (simp add: mantissa_eq_zero_iff)
63664
nipkow
parents: 63663
diff changeset
   911
  moreover
nipkow
parents: 63663
diff changeset
   912
  obtain i where "m = mantissa f * 2 ^ i" "e = exponent f - int i"
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   913
    by (rule f_def[THEN denormalize_shift, OF \<open>f \<noteq> 0\<close>])
63664
nipkow
parents: 63663
diff changeset
   914
  ultimately show ?thesis by (simp add: abs_mult)
nipkow
parents: 63663
diff changeset
   915
qed
nipkow
parents: 63663
diff changeset
   916
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   917
lemma float_gt1_scale:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
   918
  assumes "1 \<le> Float m e"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   919
  shows "0 \<le> e + (bitlen m - 1)"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   920
proof -
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   921
  have "0 < Float m e" using assms by auto
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   922
  then have "0 < m" using powr_gt_zero[of 2 e]
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
   923
    by (auto simp: zero_less_mult_iff)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   924
  then have "m \<noteq> 0" by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   925
  show ?thesis
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   926
  proof (cases "0 \<le> e")
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   927
    case True
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   928
    then show ?thesis
63248
414e3550e9c0 generalized bitlen to floor of log
immler
parents: 63040
diff changeset
   929
      using \<open>0 < m\<close> by (simp add: bitlen_alt_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   930
  next
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   931
    case False
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   932
    have "(1::int) < 2" by simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   933
    let ?S = "2^(nat (-e))"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   934
    have "inverse (2 ^ nat (- e)) = 2 powr e"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   935
      using assms False powr_realpow[of 2 "nat (-e)"]
57862
8f074e6e22fc tuned proofs;
wenzelm
parents: 57512
diff changeset
   936
      by (auto simp: powr_minus field_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   937
    then have "1 \<le> real_of_int m * inverse ?S"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   938
      using assms False powr_realpow[of 2 "nat (-e)"]
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   939
      by (auto simp: powr_minus)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   940
    then have "1 * ?S \<le> real_of_int m * inverse ?S * ?S"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   941
      by (rule mult_right_mono) auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   942
    then have "?S \<le> real_of_int m"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   943
      unfolding mult.assoc by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   944
    then have "?S \<le> m"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
   945
      unfolding of_int_le_iff[symmetric] by auto
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   946
    from this bitlen_bounds[OF \<open>0 < m\<close>, THEN conjunct2]
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   947
    have "nat (-e) < (nat (bitlen m))"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   948
      unfolding power_strict_increasing_iff[OF \<open>1 < 2\<close>, symmetric]
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   949
      by (rule order_le_less_trans)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   950
    then have "-e < bitlen m"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   951
      using False by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   952
    then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   953
      by auto
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   954
  qed
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
   955
qed
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
   956
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   957
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
   958
subsection \<open>Truncating Real Numbers\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   959
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   960
definition truncate_down::"nat \<Rightarrow> real \<Rightarrow> real"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
   961
  where "truncate_down prec x = round_down (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) x"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   962
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   963
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
   964
  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
   965
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   966
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
   967
  by (rule order_trans[OF truncate_down])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   968
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   969
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
   970
  by (simp add: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   971
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   972
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
   973
  by (auto simp: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   974
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
   975
definition truncate_up::"nat \<Rightarrow> real \<Rightarrow> real"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
   976
  where "truncate_up prec x = round_up (prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) x"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   977
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   978
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
   979
  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
   980
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   981
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
   982
  by (rule order_trans[OF _ truncate_up])
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   983
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   984
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
   985
  by (simp add: truncate_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   986
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   987
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
   988
  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
   989
  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
   990
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   991
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
   992
  by (auto simp: truncate_up_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
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
   995
  by (simp_all add: powr_add)
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_down_pos:
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
   998
  assumes "x > 0"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
   999
  shows "truncate_down p x > 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1000
proof -
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1001
  have "0 \<le> log 2 x - real_of_int \<lfloor>log 2 x\<rfloor>"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1002
    by (simp add: algebra_simps)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  1003
  with assms
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1004
  show ?thesis
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1005
    apply (auto simp: truncate_down_def round_down_def mult_powr_eq
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1006
      intro!: ge_one_powr_ge_zero mult_pos_pos)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  1007
    by linarith
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1008
qed
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_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
  1011
  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
  1012
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1013
lemma truncate_down_ge1: "1 \<le> x \<Longrightarrow> 1 \<le> truncate_down p x"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1014
  apply (auto simp: truncate_down_def algebra_simps intro!: round_down_ge1)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1015
  apply linarith
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1016
  done
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1017
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1018
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
  1019
  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
  1020
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1021
lemma truncate_up_le1:
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1022
  assumes "x \<le> 1"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1023
  shows "truncate_up p x \<le> 1"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1024
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1025
  consider "x \<le> 0" | "x > 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1026
    by arith
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1027
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1028
  proof cases
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1029
    case 1
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1030
    with truncate_up_nonpos[OF this, of p] show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1031
      by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1032
  next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1033
    case 2
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1034
    then have le: "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<le> 0"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1035
      using assms by (auto simp: log_less_iff)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1036
    from assms have "0 \<le> int p" by simp
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1037
    from add_mono[OF this le]
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1038
    show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1039
      using assms by (simp add: truncate_up_def round_up_le1 add_mono)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1040
  qed
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1041
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1042
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1043
lemma truncate_down_shift_int:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1044
  "truncate_down p (x * 2 powr real_of_int k) = truncate_down p x * 2 powr k"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1045
  by (cases "x = 0")
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1046
    (simp_all add: algebra_simps abs_mult log_mult truncate_down_def
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1047
      round_down_shift[of _ _ k, simplified])
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1048
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1049
lemma truncate_down_shift_nat: "truncate_down p (x * 2 powr real k) = truncate_down p x * 2 powr k"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1050
  by (metis of_int_of_nat_eq truncate_down_shift_int)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1051
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1052
lemma truncate_up_shift_int: "truncate_up p (x * 2 powr real_of_int k) = truncate_up p x * 2 powr k"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1053
  by (cases "x = 0")
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1054
    (simp_all add: algebra_simps abs_mult log_mult truncate_up_def
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1055
      round_up_shift[of _ _ k, simplified])
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1056
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1057
lemma truncate_up_shift_nat: "truncate_up p (x * 2 powr real k) = truncate_up p x * 2 powr k"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1058
  by (metis of_int_of_nat_eq truncate_up_shift_int)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1059
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1060
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1061
subsection \<open>Truncating Floats\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1062
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1063
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
  1064
  by (simp add: truncate_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1065
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1066
lemma float_round_up: "real_of_float x \<le> real_of_float (float_round_up prec x)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1067
  using truncate_up by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1068
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1069
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
  1070
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1071
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1072
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
  1073
  by (simp add: truncate_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1074
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1075
lemma float_round_down: "real_of_float (float_round_down prec x) \<le> real_of_float x"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1076
  using truncate_down by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1077
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1078
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
  1079
  by transfer simp
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
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
  1082
  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
  1083
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1084
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
  1085
  and minus_float_round_down_eq: "- float_round_down prec x = float_round_up prec (- x)"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1086
  by (transfer; simp add: truncate_down_uminus_eq truncate_up_uminus_eq)+
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1087
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1088
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1089
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1090
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1091
qualified lemma compute_float_round_down[code]:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1092
  "float_round_down prec (Float m e) =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1093
    (let d = bitlen \<bar>m\<bar> - int prec - 1 in
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1094
      if 0 < d then Float (div_twopow m (nat d)) (e + d)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1095
      else Float m e)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1096
  using Float.compute_float_down[of "Suc prec - bitlen \<bar>m\<bar> - e" m e, symmetric]
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1097
  by transfer
63248
414e3550e9c0 generalized bitlen to floor of log
immler
parents: 63040
diff changeset
  1098
    (simp add: field_simps abs_mult log_mult bitlen_alt_def truncate_down_def
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1099
      cong del: if_weak_cong)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1100
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1101
qualified lemma compute_float_round_up[code]:
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1102
  "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
  1103
  by transfer (simp add: truncate_down_uminus_eq)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1104
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1105
end
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1106
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1107
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1108
subsection \<open>Approximation of positive rationals\<close>
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1109
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1110
lemma div_mult_twopow_eq: "a div ((2::nat) ^ n) div b = a div (b * 2 ^ n)" for a b :: nat
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1111
  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
  1112
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1113
lemma real_div_nat_eq_floor_of_divide: "a div b = real_of_int \<lfloor>a / b\<rfloor>" for a b :: nat
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1114
  by (simp add: floor_divide_of_nat_eq [of a b])
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1115
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1116
definition "rat_precision prec x y =
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1117
  (let d = bitlen x - bitlen y
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1118
   in int prec - d + (if Float (abs x) 0 < Float (abs y) d then 1 else 0))"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1119
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1120
lemma floor_log_divide_eq:
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1121
  assumes "i > 0" "j > 0" "p > 1"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1122
  shows "\<lfloor>log p (i / j)\<rfloor> = floor (log p i) - floor (log p j) -
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1123
    (if i \<ge> j * p powr (floor (log p i) - floor (log p j)) then 0 else 1)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1124
proof -
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1125
  let ?l = "log p"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1126
  let ?fl = "\<lambda>x. floor (?l x)"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1127
  have "\<lfloor>?l (i / j)\<rfloor> = \<lfloor>?l i - ?l j\<rfloor>" using assms
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1128
    by (auto simp: log_divide)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1129
  also have "\<dots> = floor (real_of_int (?fl i - ?fl j) + (?l i - ?fl i - (?l j - ?fl j)))"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1130
    (is "_ = floor (_ + ?r)")
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1131
    by (simp add: algebra_simps)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1132
  also note floor_add2
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1133
  also note \<open>p > 1\<close>
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1134
  note powr = powr_le_cancel_iff[symmetric, OF \<open>1 < p\<close>, THEN iffD2]
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1135
  note powr_strict = powr_less_cancel_iff[symmetric, OF \<open>1 < p\<close>, THEN iffD2]
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1136
  have "floor ?r = (if i \<ge> j * p powr (?fl i - ?fl j) then 0 else -1)" (is "_ = ?if")
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1137
    using assms
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1138
    by (linarith |
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1139
      auto
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1140
        intro!: floor_eq2
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1141
        intro: powr_strict powr
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1142
        simp: powr_diff powr_add divide_simps algebra_simps)+
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1143
  finally
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1144
  show ?thesis by simp
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1145
qed
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1146
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1147
lemma truncate_down_rat_precision:
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1148
    "truncate_down prec (real x / real y) = round_down (rat_precision prec x y) (real x / real y)"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1149
  and truncate_up_rat_precision:
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1150
    "truncate_up prec (real x / real y) = round_up (rat_precision prec x y) (real x / real y)"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1151
  unfolding truncate_down_def truncate_up_def rat_precision_def
63248
414e3550e9c0 generalized bitlen to floor of log
immler
parents: 63040
diff changeset
  1152
  by (cases x; cases y) (auto simp: floor_log_divide_eq algebra_simps bitlen_alt_def)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1153
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1154
lift_definition lapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1155
  is "\<lambda>prec (x::nat) (y::nat). truncate_down prec (x / y)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1156
  by simp
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1157
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1158
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1159
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1160
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1161
qualified lemma compute_lapprox_posrat[code]:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1162
  "lapprox_posrat prec x y =
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1163
   (let
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1164
      l = rat_precision prec x y;
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1165
      d = if 0 \<le> l then x * 2^nat l div y else x div 2^nat (- l) div y
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1166
    in normfloat (Float d (- l)))"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1167
    unfolding div_mult_twopow_eq
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1168
    by transfer
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1169
      (simp add: round_down_def powr_int real_div_nat_eq_floor_of_divide field_simps Let_def
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1170
        truncate_down_rat_precision del: two_powr_minus_int_float)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1171
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1172
end
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1173
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1174
lift_definition rapprox_posrat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1175
  is "\<lambda>prec (x::nat) (y::nat). truncate_up prec (x / y)"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1176
  by simp
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1177
60376
wenzelm
parents: 60017
diff changeset
  1178
context
wenzelm
parents: 60017
diff changeset
  1179
begin
wenzelm
parents: 60017
diff changeset
  1180
wenzelm
parents: 60017
diff changeset
  1181
qualified lemma compute_rapprox_posrat[code]:
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1182
  fixes prec x y
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1183
  defines "l \<equiv> rat_precision prec x y"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1184
  shows "rapprox_posrat prec x y =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1185
   (let
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1186
      l = l;
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1187
      (r, s) = if 0 \<le> l then (x * 2^nat l, y) else (x, y * 2^nat(-l));
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1188
      d = r div s;
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1189
      m = r mod s
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1190
    in normfloat (Float (d + (if m = 0 \<or> y = 0 then 0 else 1)) (- l)))"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1191
proof (cases "y = 0")
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1192
  assume "y = 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1193
  then show ?thesis by transfer simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1194
next
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1195
  assume "y \<noteq> 0"
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1196
  show ?thesis
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1197
  proof (cases "0 \<le> l")
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1198
    case True
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62421
diff changeset
  1199
    define x' where "x' = x * 2 ^ nat l"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1200
    have "int x * 2 ^ nat l = x'"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1201
      by (simp add: x'_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1202
    moreover have "real x * 2 powr l = real x'"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1203
      by (simp flip: powr_realpow add: \<open>0 \<le> l\<close> x'_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1204
    ultimately show ?thesis
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1205
      using ceil_divide_floor_conv[of y x'] powr_realpow[of 2 "nat l"] \<open>0 \<le> l\<close> \<open>y \<noteq> 0\<close>
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1206
        l_def[symmetric, THEN meta_eq_to_obj_eq]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1207
      apply transfer
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1208
      apply (auto simp add: round_up_def truncate_up_rat_precision)
67118
ccab07d1196c more simplification rules
haftmann
parents: 66912
diff changeset
  1209
      apply (metis dvd_triv_left of_nat_dvd_iff)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1210
      apply (metis floor_divide_of_int_eq of_int_of_nat_eq)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1211
      done
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1212
   next
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1213
    case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62421
diff changeset
  1214
    define y' where "y' = y * 2 ^ nat (- l)"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1215
    from \<open>y \<noteq> 0\<close> have "y' \<noteq> 0" by (simp add: y'_def)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1216
    have "int y * 2 ^ nat (- l) = y'"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1217
      by (simp add: y'_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1218
    moreover have "real x * real_of_int (2::int) powr real_of_int l / real y = x / real y'"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1219
      using \<open>\<not> 0 \<le> l\<close> by (simp flip: powr_realpow add: powr_minus y'_def field_simps)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1220
    ultimately show ?thesis
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1221
      using ceil_divide_floor_conv[of y' x] \<open>\<not> 0 \<le> l\<close> \<open>y' \<noteq> 0\<close> \<open>y \<noteq> 0\<close>
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1222
        l_def[symmetric, THEN meta_eq_to_obj_eq]
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1223
      apply transfer
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1224
      apply (auto simp add: round_up_def ceil_divide_floor_conv truncate_up_rat_precision)
67118
ccab07d1196c more simplification rules
haftmann
parents: 66912
diff changeset
  1225
      apply (metis dvd_triv_left of_nat_dvd_iff)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1226
      apply (metis floor_divide_of_int_eq of_int_of_nat_eq)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1227
      done
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1228
  qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1229
qed
60376
wenzelm
parents: 60017
diff changeset
  1230
wenzelm
parents: 60017
diff changeset
  1231
end
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1232
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1233
lemma rat_precision_pos:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1234
  assumes "0 \<le> x"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1235
    and "0 < y"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1236
    and "2 * x < y"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1237
  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
  1238
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1239
  have "0 < x \<Longrightarrow> log 2 x + 1 = log 2 (2 * x)"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1240
    by (simp add: log_mult)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1241
  then have "bitlen (int x) < bitlen (int y)"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1242
    using assms
63599
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1243
    by (simp add: bitlen_alt_def)
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1244
      (auto intro!: floor_mono simp add: one_add_floor)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1245
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1246
    using assms
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1247
    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
  1248
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1249
47601
050718fe6eee use real :: float => real as lifting-morphism so we can directlry use the rep_eq theorems
hoelzl
parents: 47600
diff changeset
  1250
lemma rapprox_posrat_less1:
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1251
  "0 \<le> x \<Longrightarrow> 0 < y \<Longrightarrow> 2 * x < y \<Longrightarrow> real_of_float (rapprox_posrat n x y) < 1"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1252
  by transfer (simp add: rat_precision_pos round_up_less1 truncate_up_rat_precision)
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1253
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1254
lift_definition lapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1255
  "\<lambda>prec (x::int) (y::int). truncate_down prec (x / y)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1256
  by simp
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1257
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1258
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1259
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1260
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1261
qualified lemma compute_lapprox_rat[code]:
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1262
  "lapprox_rat prec x y =
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1263
   (if y = 0 then 0
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1264
    else if 0 \<le> x then
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1265
     (if 0 < y then lapprox_posrat prec (nat x) (nat y)
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  1266
      else - (rapprox_posrat prec (nat x) (nat (-y))))
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1267
      else
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1268
       (if 0 < y
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1269
        then - (rapprox_posrat prec (nat (-x)) (nat y))
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1270
        else lapprox_posrat prec (nat (-x)) (nat (-y))))"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1271
  by transfer (simp add: truncate_up_uminus_eq)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1272
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1273
lift_definition rapprox_rat :: "nat \<Rightarrow> int \<Rightarrow> int \<Rightarrow> float" is
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1274
  "\<lambda>prec (x::int) (y::int). truncate_up prec (x / y)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1275
  by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1276
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1277
lemma "rapprox_rat = rapprox_posrat"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1278
  by transfer auto
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1279
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1280
lemma "lapprox_rat = lapprox_posrat"
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1281
  by transfer auto
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1282
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1283
qualified lemma compute_rapprox_rat[code]:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1284
  "rapprox_rat prec x y = - lapprox_rat prec (-x) y"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1285
  by transfer (simp add: truncate_down_uminus_eq)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1286
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1287
qualified lemma compute_truncate_down[code]:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1288
  "truncate_down p (Ratreal r) = (let (a, b) = quotient_of r in lapprox_rat p a b)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1289
  by transfer (auto split: prod.split simp: of_rat_divide dest!: quotient_of_div)
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1290
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1291
qualified lemma compute_truncate_up[code]:
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1292
  "truncate_up p (Ratreal r) = (let (a, b) = quotient_of r in rapprox_rat p a b)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1293
  by transfer (auto split: prod.split simp:  of_rat_divide dest!: quotient_of_div)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1294
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1295
end
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1296
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1297
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1298
subsection \<open>Division\<close>
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1299
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1300
definition "real_divl prec a b = truncate_down prec (a / b)"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1301
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1302
definition "real_divr prec a b = truncate_up prec (a / b)"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1303
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1304
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
  1305
  by (simp add: real_divl_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1306
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1307
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1308
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1309
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1310
qualified lemma compute_float_divl[code]:
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1311
  "float_divl prec (Float m1 s1) (Float m2 s2) = lapprox_rat prec m1 m2 * Float 1 (s1 - s2)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1312
  apply transfer
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1313
  unfolding real_divl_def of_int_1 mult_1 truncate_down_shift_int[symmetric]
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1314
  apply (simp add: powr_diff powr_minus)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1315
  done
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1316
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1317
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
  1318
  by (simp add: real_divr_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1319
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1320
qualified lemma compute_float_divr[code]:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1321
  "float_divr prec x y = - float_divl prec (-x) y"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1322
  by transfer (simp add: real_divr_def real_divl_def truncate_down_uminus_eq)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1323
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1324
end
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  1325
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1326
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1327
subsection \<open>Approximate Power\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1328
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1329
lemma div2_less_self[termination_simp]: "odd n \<Longrightarrow> n div 2 < n" for n :: nat
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1330
  by (simp add: odd_pos)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1331
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1332
fun power_down :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1333
where
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1334
  "power_down p x 0 = 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1335
| "power_down p x (Suc n) =
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1336
    (if odd n then truncate_down (Suc p) ((power_down p x (Suc n div 2))\<^sup>2)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1337
     else truncate_down (Suc p) (x * power_down p x n))"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1338
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1339
fun power_up :: "nat \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1340
where
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1341
  "power_up p x 0 = 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1342
| "power_up p x (Suc n) =
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1343
    (if odd n then truncate_up p ((power_up p x (Suc n div 2))\<^sup>2)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1344
     else truncate_up p (x * power_up p x n))"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1345
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1346
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
  1347
  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
  1348
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1349
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
  1350
  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
  1351
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1352
lemma power_float_transfer[transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67118
diff changeset
  1353
  "(rel_fun pcr_float (rel_fun (=) pcr_float)) (^) (^)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1354
  unfolding power_def
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1355
  by transfer_prover
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1356
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1357
lemma compute_power_up_fl[code]:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1358
    "power_up_fl p x 0 = 1"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1359
    "power_up_fl p x (Suc n) =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1360
      (if odd n then float_round_up p ((power_up_fl p x (Suc n div 2))\<^sup>2)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1361
       else float_round_up p (x * power_up_fl p x n))"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1362
  and compute_power_down_fl[code]:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1363
    "power_down_fl p x 0 = 1"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1364
    "power_down_fl p x (Suc n) =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1365
      (if odd n then float_round_down (Suc p) ((power_down_fl p x (Suc n div 2))\<^sup>2)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1366
       else float_round_down (Suc p) (x * power_down_fl p x n))"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1367
  unfolding atomize_conj by transfer simp
58985
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 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
  1370
  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
  1371
    (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
  1372
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1373
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
  1374
  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
  1375
    (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
  1376
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1377
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
  1378
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
  1379
  case (2 p x n)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1380
  have ?case if "odd n"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1381
  proof -
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1382
    from that 2 have "(power_down p x (Suc n div 2)) ^ 2 \<le> (x ^ (Suc n div 2)) ^ 2"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1383
      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
  1384
    also have "\<dots> = x ^ (Suc n div 2 * 2)"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1385
      by (simp flip: power_mult)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1386
    also have "Suc n div 2 * 2 = Suc n"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1387
      using \<open>odd n\<close> by presburger
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1388
    finally show ?thesis
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1389
      using that by (auto intro!: truncate_down_le simp del: odd_Suc_div_two)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1390
  qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1391
  then show ?case
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1392
    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
  1393
qed simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1394
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1395
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
  1396
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
  1397
  case (2 p x n)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1398
  have ?case if "odd n"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1399
  proof -
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1400
    from that even_Suc have "Suc n = Suc n div 2 * 2"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1401
      by presburger
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1402
    then have "x ^ Suc n \<le> (x ^ (Suc n div 2))\<^sup>2"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1403
      by (simp flip: power_mult)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1404
    also from that 2 have "\<dots> \<le> (power_up p x (Suc n div 2))\<^sup>2"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1405
      by (auto intro: power_mono simp del: odd_Suc_div_two)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1406
    finally show ?thesis
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1407
      using that by (auto intro!: truncate_up_le simp del: odd_Suc_div_two)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1408
  qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1409
  then show ?case
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1410
    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
  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
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
  1414
  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
  1415
  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
  1416
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1417
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
  1418
  by transfer (rule power_down)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1419
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1420
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
  1421
  by transfer (rule power_up)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1422
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1423
lemma real_power_up_fl: "real_of_float (power_up_fl p x n) = power_up p x n"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1424
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1425
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1426
lemma real_power_down_fl: "real_of_float (power_down_fl p x n) = power_down p x n"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1427
  by transfer simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1428
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1429
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1430
subsection \<open>Approximate Addition\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1431
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1432
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
  1433
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1434
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
  1435
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1436
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
  1437
  by (simp add: plus_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1438
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1439
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
  1440
  by (simp add: plus_up_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1441
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1442
lift_definition float_plus_down :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is plus_down ..
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1443
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1444
lift_definition float_plus_up :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float" is plus_up ..
58985
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
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
  1447
  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
  1448
  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
  1449
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1450
lemma float_plus_down: "real_of_float (float_plus_down prec x y) \<le> x + y"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1451
  and float_plus_up: "x + y \<le> real_of_float (float_plus_up prec x y)"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1452
  by (transfer; rule plus_down plus_up)+
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1453
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1454
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
  1455
  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
  1456
  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
  1457
  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
  1458
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1459
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
  1460
  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
  1461
  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
  1462
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1463
lemma truncate_down_log2_eqI:
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1464
  assumes "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1465
  assumes "\<lfloor>x * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)\<rfloor> = \<lfloor>y * 2 powr (p - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)\<rfloor>"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1466
  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
  1467
  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
  1468
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1469
lemma sum_neq_zeroI:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1470
  "\<bar>a\<bar> \<ge> k \<Longrightarrow> \<bar>b\<bar> < k \<Longrightarrow> a + b \<noteq> 0"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1471
  "\<bar>a\<bar> > k \<Longrightarrow> \<bar>b\<bar> \<le> k \<Longrightarrow> a + b \<noteq> 0"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1472
  for a k :: real
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1473
  by auto
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1474
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1475
lemma abs_real_le_2_powr_bitlen[simp]: "\<bar>real_of_int m2\<bar> < 2 powr real_of_int (bitlen \<bar>m2\<bar>)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1476
proof (cases "m2 = 0")
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1477
  case True
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1478
  then show ?thesis by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1479
next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1480
  case False
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1481
  then have "\<bar>m2\<bar> < 2 ^ nat (bitlen \<bar>m2\<bar>)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1482
    using bitlen_bounds[of "\<bar>m2\<bar>"]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1483
    by (auto simp: powr_add bitlen_nonneg)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1484
  then show ?thesis
66912
a99a7cbf0fb5 generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents: 65583
diff changeset
  1485
    by (metis bitlen_nonneg powr_int of_int_abs of_int_less_numeral_power_cancel_iff
a99a7cbf0fb5 generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents: 65583
diff changeset
  1486
        zero_less_numeral)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1487
qed
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1488
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1489
lemma floor_sum_times_2_powr_sgn_eq:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1490
  fixes ai p q :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1491
    and a b :: real
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1492
  assumes "a * 2 powr p = ai"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1493
    and b_le_1: "\<bar>b * 2 powr (p + 1)\<bar> \<le> 1"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1494
    and leqp: "q \<le> p"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1495
  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
  1496
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1497
  consider "b = 0" | "b > 0" | "b < 0" by arith
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1498
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1499
  proof cases
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1500
    case 1
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1501
    then show ?thesis
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1502
      by (simp flip: assms(1) powr_add add: algebra_simps powr_mult_base)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1503
  next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1504
    case 2
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1505
    then have "b * 2 powr p < \<bar>b * 2 powr (p + 1)\<bar>"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1506
      by simp
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1507
    also note b_le_1
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1508
    finally have b_less_1: "b * 2 powr real_of_int p < 1" .
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1509
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1510
    from b_less_1 \<open>b > 0\<close> have floor_eq: "\<lfloor>b * 2 powr real_of_int p\<rfloor> = 0" "\<lfloor>sgn b / 2\<rfloor> = 0"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1511
      by (simp_all add: floor_eq_iff)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1512
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1513
    have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(a + b) * 2 powr p * 2 powr (q - p)\<rfloor>"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1514
      by (simp add: algebra_simps flip: powr_realpow powr_add)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1515
    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
  1516
      by (simp add: assms algebra_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1517
    also have "\<dots> = \<lfloor>(ai + b * 2 powr p) / real_of_int ((2::int) ^ nat (p - q))\<rfloor>"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1518
      using assms
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1519
      by (simp add: algebra_simps divide_powr_uminus flip: powr_realpow powr_add)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1520
    also have "\<dots> = \<lfloor>ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1521
      by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1522
    finally have "\<lfloor>(a + b) * 2 powr real_of_int q\<rfloor> = \<lfloor>real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>" .
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1523
    moreover
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1524
    have "\<lfloor>(2 * ai + (sgn b)) * 2 powr (real_of_int (q - p) - 1)\<rfloor> =
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1525
        \<lfloor>real_of_int ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1526
    proof -
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1527
      have "\<lfloor>(2 * ai + sgn b) * 2 powr (real_of_int (q - p) - 1)\<rfloor> = \<lfloor>(ai + sgn b / 2) * 2 powr (q - p)\<rfloor>"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1528
        by (subst powr_diff) (simp add: field_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1529
      also have "\<dots> = \<lfloor>(ai + sgn b / 2) / real_of_int ((2::int) ^ nat (p - q))\<rfloor>"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1530
        using leqp by (simp flip: powr_realpow add: powr_diff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1531
      also have "\<dots> = \<lfloor>ai / real_of_int ((2::int) ^ nat (p - q))\<rfloor>"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1532
        by (simp del: of_int_power add: floor_divide_real_eq_div floor_eq)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1533
      finally show ?thesis .
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1534
    qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1535
    ultimately show ?thesis by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1536
  next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1537
    case 3
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1538
    then have floor_eq: "\<lfloor>b * 2 powr (real_of_int p + 1)\<rfloor> = -1"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1539
      using b_le_1
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1540
      by (auto simp: floor_eq_iff algebra_simps pos_divide_le_eq[symmetric] abs_if divide_powr_uminus
62390
842917225d56 more canonical names
nipkow
parents: 62348
diff changeset
  1541
        intro!: mult_neg_pos split: if_split_asm)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1542
    have "\<lfloor>(a + b) * 2 powr q\<rfloor> = \<lfloor>(2*a + 2*b) * 2 powr p * 2 powr (q - p - 1)\<rfloor>"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1543
      by (simp add: algebra_simps powr_mult_base flip: powr_realpow powr_add)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1544
    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
  1545
      by (simp add: algebra_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1546
    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
  1547
      using assms by (simp add: algebra_simps powr_mult_base divide_powr_uminus)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1548
    also have "\<dots> = \<lfloor>(2 * ai + b * 2 powr (p + 1)) / real_of_int ((2::int) ^ nat (p - q + 1))\<rfloor>"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1549
      using assms by (simp add: algebra_simps flip: powr_realpow)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1550
    also have "\<dots> = \<lfloor>(2 * ai - 1) / real_of_int ((2::int) ^ nat (p - q + 1))\<rfloor>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1551
      using \<open>b < 0\<close> assms
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1552
      by (simp add: floor_divide_of_int_eq floor_eq floor_divide_real_eq_div
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1553
        del: of_int_mult of_int_power of_int_diff)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1554
    also have "\<dots> = \<lfloor>(2 * ai - 1) * 2 powr (q - p - 1)\<rfloor>"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1555
      using assms by (simp add: algebra_simps divide_powr_uminus flip: powr_realpow)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1556
    finally show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1557
      using \<open>b < 0\<close> by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1558
  qed
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1559
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1560
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1561
lemma log2_abs_int_add_less_half_sgn_eq:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1562
  fixes ai :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1563
    and b :: real
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1564
  assumes "\<bar>b\<bar> \<le> 1/2"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1565
    and "ai \<noteq> 0"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1566
  shows "\<lfloor>log 2 \<bar>real_of_int ai + b\<bar>\<rfloor> = \<lfloor>log 2 \<bar>ai + sgn b / 2\<bar>\<rfloor>"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1567
proof (cases "b = 0")
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1568
  case True
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1569
  then show ?thesis by simp
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1570
next
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1571
  case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62421
diff changeset
  1572
  define k where "k = \<lfloor>log 2 \<bar>ai\<bar>\<rfloor>"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1573
  then have "\<lfloor>log 2 \<bar>ai\<bar>\<rfloor> = k"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1574
    by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1575
  then have k: "2 powr k \<le> \<bar>ai\<bar>" "\<bar>ai\<bar> < 2 powr (k + 1)"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1576
    by (simp_all add: floor_log_eq_powr_iff \<open>ai \<noteq> 0\<close>)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1577
  have "k \<ge> 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1578
    using assms by (auto simp: k_def)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62421
diff changeset
  1579
  define r where "r = \<bar>ai\<bar> - 2 ^ nat k"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1580
  have r: "0 \<le> r" "r < 2 powr k"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1581
    using \<open>k \<ge> 0\<close> k
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1582
    by (auto simp: r_def k_def algebra_simps powr_add abs_if powr_int)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1583
  then have "r \<le> (2::int) ^ nat k - 1"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1584
    using \<open>k \<ge> 0\<close> by (auto simp: powr_int)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1585
  from this[simplified of_int_le_iff[symmetric]] \<open>0 \<le> k\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1586
  have r_le: "r \<le> 2 powr k - 1"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1587
    by (auto simp: algebra_simps powr_int)
66912
a99a7cbf0fb5 generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents: 65583
diff changeset
  1588
      (metis of_int_1 of_int_add of_int_le_numeral_power_cancel_iff)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1589
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1590
  have "\<bar>ai\<bar> = 2 powr k + r"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1591
    using \<open>k \<ge> 0\<close> by (auto simp: k_def r_def simp flip: powr_realpow)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1592
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1593
  have pos: "\<bar>b\<bar> < 1 \<Longrightarrow> 0 < 2 powr k + (r + b)" for b :: real
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1594
    using \<open>0 \<le> k\<close> \<open>ai \<noteq> 0\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1595
    by (auto simp add: r_def powr_realpow[symmetric] abs_if sgn_if algebra_simps
62390
842917225d56 more canonical names
nipkow
parents: 62348
diff changeset
  1596
      split: if_split_asm)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1597
  have less: "\<bar>sgn ai * b\<bar> < 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1598
    and less': "\<bar>sgn (sgn ai * b) / 2\<bar> < 1"
62390
842917225d56 more canonical names
nipkow
parents: 62348
diff changeset
  1599
    using \<open>\<bar>b\<bar> \<le> _\<close> by (auto simp: abs_if sgn_if split: if_split_asm)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1600
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1601
  have floor_eq: "\<And>b::real. \<bar>b\<bar> \<le> 1 / 2 \<Longrightarrow>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1602
      \<lfloor>log 2 (1 + (r + b) / 2 powr k)\<rfloor> = (if r = 0 \<and> b < 0 then -1 else 0)"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1603
    using \<open>k \<ge> 0\<close> r r_le
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1604
    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
  1605
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1606
  from \<open>real_of_int \<bar>ai\<bar> = _\<close> have "\<bar>ai + b\<bar> = 2 powr k + (r + sgn ai * b)"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1607
    using \<open>\<bar>b\<bar> \<le> _\<close> \<open>0 \<le> k\<close> r
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1608
    by (auto simp add: sgn_if abs_if)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1609
  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
  1610
  proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1611
    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
  1612
      by (simp add: field_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1613
    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
  1614
      using pos[OF less]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1615
      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
  1616
    also
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1617
    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
  1618
    have "\<lfloor>log 2 (1 + (r + sgn ai * b) / 2 powr k)\<rfloor> = ?if"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1619
      using \<open>\<bar>b\<bar> \<le> _\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1620
      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
  1621
    also
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1622
    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
  1623
      by (subst floor_eq) (auto simp: sgn_if)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1624
    also have "k + \<dots> = \<lfloor>log 2 (2 powr k * (1 + (r + sgn (sgn ai * b) / 2) / 2 powr k))\<rfloor>"
63599
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1625
      unfolding int_add_floor
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1626
      using pos[OF less'] \<open>\<bar>b\<bar> \<le> _\<close>
63599
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1627
      by (simp add: field_simps add_log_eq_powr del: floor_add2)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1628
    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
  1629
        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
  1630
      by (simp add: sgn_if field_simps)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1631
    finally show ?thesis .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1632
  qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1633
  also have "2 powr k + r + sgn (sgn ai * b) / 2 = \<bar>ai + sgn b / 2\<bar>"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1634
    unfolding \<open>real_of_int \<bar>ai\<bar> = _\<close>[symmetric] using \<open>ai \<noteq> 0\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1635
    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
  1636
  finally show ?thesis .
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1637
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1638
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1639
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1640
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1641
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1642
qualified lemma compute_far_float_plus_down:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1643
  fixes m1 e1 m2 e2 :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1644
    and p :: nat
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1645
  defines "k1 \<equiv> Suc p - nat (bitlen \<bar>m1\<bar>)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1646
  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
  1647
  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
  1648
    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
  1649
proof -
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1650
  let ?a = "real_of_float (Float m1 e1)"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1651
  let ?b = "real_of_float (Float m2 e2)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1652
  let ?sum = "?a + ?b"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1653
  let ?shift = "real_of_int e2 - real_of_int e1 + real k1 + 1"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1654
  let ?m1 = "m1 * 2 ^ Suc k1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1655
  let ?m2 = "m2 * 2 powr ?shift"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1656
  let ?m2' = "sgn m2 / 2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1657
  let ?e = "e1 - int k1 - 1"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1658
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1659
  have sum_eq: "?sum = (?m1 + ?m2) * 2 powr ?e"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1660
    by (auto simp flip: powr_add powr_mult powr_realpow simp: powr_mult_base algebra_simps)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1661
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1662
  have "\<bar>?m2\<bar> * 2 < 2 powr (bitlen \<bar>m2\<bar> + ?shift + 1)"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1663
    by (auto simp: field_simps powr_add powr_mult_base powr_diff abs_mult)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1664
  also have "\<dots> \<le> 2 powr 0"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1665
    using H by (intro powr_mono) auto
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1666
  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
  1667
    by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1668
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1669
  then have "\<bar>real_of_int m2\<bar> < 2 powr -(?shift + 1)"
63248
414e3550e9c0 generalized bitlen to floor of log
immler
parents: 63040
diff changeset
  1670
    unfolding powr_minus_divide by (auto simp: bitlen_alt_def field_simps powr_mult_base abs_mult)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1671
  also have "\<dots> \<le> 2 powr real_of_int (e1 - e2 - 2)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1672
    by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1673
  finally have b_less_quarter: "\<bar>?b\<bar> < 1/4 * 2 powr real_of_int e1"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1674
    by (simp add: powr_add field_simps powr_diff abs_mult)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1675
  also have "1/4 < \<bar>real_of_int m1\<bar> / 2" using \<open>m1 \<noteq> 0\<close> by simp
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1676
  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
  1677
    by (simp add: algebra_simps powr_mult_base abs_mult)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1678
  then have a_half_less_sum: "\<bar>?a\<bar> / 2 < \<bar>?sum\<bar>"
62390
842917225d56 more canonical names
nipkow
parents: 62348
diff changeset
  1679
    by (auto simp: field_simps abs_if split: if_split_asm)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1680
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1681
  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
  1682
    by simp_all
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1683
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1684
  have "\<bar>real_of_float (Float m1 e1)\<bar> \<ge> 1/4 * 2 powr real_of_int e1"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1685
    using \<open>m1 \<noteq> 0\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1686
    by (auto simp: powr_add powr_int bitlen_nonneg divide_right_mono abs_mult)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1687
  then have "?sum \<noteq> 0" using b_less_quarter
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1688
    by (rule sum_neq_zeroI)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1689
  then have "?m1 + ?m2 \<noteq> 0"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1690
    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
  1691
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1692
  have "\<bar>real_of_int ?m1\<bar> \<ge> 2 ^ Suc k1" "\<bar>?m2'\<bar> < 2 ^ Suc k1"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1693
    using \<open>m1 \<noteq> 0\<close> \<open>m2 \<noteq> 0\<close> by (auto simp: sgn_if less_1_mult abs_mult simp del: power.simps)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1694
  then have sum'_nz: "?m1 + ?m2' \<noteq> 0"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1695
    by (intro sum_neq_zeroI)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1696
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1697
  have "\<lfloor>log 2 \<bar>real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> + ?e"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1698
    using \<open>?m1 + ?m2 \<noteq> 0\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1699
    unfolding floor_add[symmetric] sum_eq
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1700
    by (simp add: abs_mult log_mult) linarith
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1701
  also have "\<lfloor>log 2 \<bar>?m1 + ?m2\<bar>\<rfloor> = \<lfloor>log 2 \<bar>?m1 + sgn (real_of_int m2 * 2 powr ?shift) / 2\<bar>\<rfloor>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1702
    using abs_m2_less_half \<open>m1 \<noteq> 0\<close>
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1703
    by (intro log2_abs_int_add_less_half_sgn_eq) (auto simp: abs_mult)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1704
  also have "sgn (real_of_int m2 * 2 powr ?shift) = sgn m2"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1705
    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
  1706
  also
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1707
  have "\<bar>?m1 + ?m2'\<bar> * 2 powr ?e = \<bar>?m1 * 2 + sgn m2\<bar> * 2 powr (?e - 1)"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1708
    by (auto simp: field_simps powr_minus[symmetric] powr_diff powr_mult_base)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1709
  then have "\<lfloor>log 2 \<bar>?m1 + ?m2'\<bar>\<rfloor> + ?e = \<lfloor>log 2 \<bar>real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1))\<bar>\<rfloor>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1710
    using \<open>?m1 + ?m2' \<noteq> 0\<close>
63599
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1711
    unfolding floor_add_int
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1712
    by (simp add: log_add_eq_powr abs_mult_pos del: floor_add2)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1713
  finally
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1714
  have "\<lfloor>log 2 \<bar>?sum\<bar>\<rfloor> = \<lfloor>log 2 \<bar>real_of_float (Float (?m1*2 + sgn m2) (?e - 1))\<bar>\<rfloor>" .
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1715
  then have "plus_down p (Float m1 e1) (Float m2 e2) =
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1716
      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
  1717
    unfolding plus_down_def
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1718
  proof (rule truncate_down_log2_eqI)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1719
    let ?f = "(int p - \<lfloor>log 2 \<bar>real_of_float (Float m1 e1) + real_of_float (Float m2 e2)\<bar>\<rfloor>)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1720
    let ?ai = "m1 * 2 ^ (Suc k1)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1721
    have "\<lfloor>(?a + ?b) * 2 powr real_of_int ?f\<rfloor> = \<lfloor>(real_of_int (2 * ?ai) + sgn ?b) * 2 powr real_of_int (?f - - ?e - 1)\<rfloor>"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1722
    proof (rule floor_sum_times_2_powr_sgn_eq)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1723
      show "?a * 2 powr real_of_int (-?e) = real_of_int ?ai"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1724
        by (simp add: powr_add powr_realpow[symmetric] powr_diff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1725
      show "\<bar>?b * 2 powr real_of_int (-?e + 1)\<bar> \<le> 1"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1726
        using abs_m2_less_half
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1727
        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
  1728
    next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1729
      have "e1 + \<lfloor>log 2 \<bar>real_of_int m1\<bar>\<rfloor> - 1 = \<lfloor>log 2 \<bar>?a\<bar>\<rfloor> - 1"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1730
        using \<open>m1 \<noteq> 0\<close>
63599
f560147710fb fixed floor proofs
nipkow
parents: 63596
diff changeset
  1731
        by (simp add: int_add_floor algebra_simps log_mult abs_mult del: floor_add2)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1732
      also have "\<dots> \<le> \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1733
        using a_half_less_sum \<open>m1 \<noteq> 0\<close> \<open>?sum \<noteq> 0\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1734
        unfolding floor_diff_of_int[symmetric]
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1735
        by (auto simp add: log_minus_eq_powr powr_minus_divide intro!: floor_mono)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1736
      finally
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1737
      have "int p - \<lfloor>log 2 \<bar>?a + ?b\<bar>\<rfloor> \<le> p - (bitlen \<bar>m1\<bar>) - e1 + 2"
63248
414e3550e9c0 generalized bitlen to floor of log
immler
parents: 63040
diff changeset
  1738
        by (auto simp: algebra_simps bitlen_alt_def \<open>m1 \<noteq> 0\<close>)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1739
      also have "\<dots> \<le> - ?e"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1740
        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
  1741
      finally show "?f \<le> - ?e" by simp
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1742
    qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1743
    also have "sgn ?b = sgn m2"
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1744
      using powr_gt_zero[of 2 e2]
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1745
      by (auto simp add: sgn_if zero_less_mult_iff simp del: powr_gt_zero)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1746
    also have "\<lfloor>(real_of_int (2 * ?m1) + real_of_int (sgn m2)) * 2 powr real_of_int (?f - - ?e - 1)\<rfloor> =
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1747
        \<lfloor>Float (?m1 * 2 + sgn m2) (?e - 1) * 2 powr ?f\<rfloor>"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1748
      by (simp flip: powr_add powr_realpow add: algebra_simps)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1749
    finally
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1750
    show "\<lfloor>(?a + ?b) * 2 powr ?f\<rfloor> = \<lfloor>real_of_float (Float (?m1 * 2 + sgn m2) (?e - 1)) * 2 powr ?f\<rfloor>" .
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1751
  qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1752
  then show ?thesis
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1753
    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
  1754
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1755
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1756
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
  1757
  by transfer (auto simp: plus_down_def)
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1758
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1759
qualified lemma compute_float_plus_down[code]:
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1760
  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
  1761
  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
  1762
    (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
  1763
    else if m2 = 0 then float_round_down p (Float m1 e1)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1764
    else
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1765
      (if e1 \<ge> e2 then
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1766
        (let k1 = Suc p - nat (bitlen \<bar>m1\<bar>) in
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1767
          if bitlen \<bar>m2\<bar> > e1 - e2 - k1 - 2
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1768
          then float_round_down p ((Float m1 e1) + (Float m2 e2))
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1769
          else float_round_down p (Float (m1 * 2 ^ (Suc (Suc k1)) + sgn m2) (e1 - int k1 - 2)))
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1770
    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
  1771
proof -
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1772
  {
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1773
    assume "bitlen \<bar>m2\<bar> \<le> e1 - e2 - (Suc p - nat (bitlen \<bar>m1\<bar>)) - 2" "m1 \<noteq> 0" "m2 \<noteq> 0" "e1 \<ge> e2"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1774
    note compute_far_float_plus_down[OF this]
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1775
  }
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1776
  then show ?thesis
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1777
    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
  1778
qed
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1779
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1780
qualified lemma compute_float_plus_up[code]: "float_plus_up p x y = - float_plus_down p (-x) (-y)"
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1781
  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
  1782
  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
  1783
70347
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
  1784
lemma mantissa_zero: "mantissa 0 = 0"
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
  1785
  by (fact mantissa_0)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1786
62421
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1787
qualified lemma compute_float_less[code]: "a < b \<longleftrightarrow> is_float_pos (float_plus_down 0 b (- a))"
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1788
  using truncate_down[of 0 "b - a"] truncate_down_pos[of "b - a" 0]
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1789
  by transfer (auto simp: plus_down_def)
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1790
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1791
qualified lemma compute_float_le[code]: "a \<le> b \<longleftrightarrow> is_float_nonneg (float_plus_down 0 b (- a))"
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1792
  using truncate_down[of 0 "b - a"] truncate_down_nonneg[of "b - a" 0]
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1793
  by transfer (auto simp: plus_down_def)
28d2c75dd180 finite precision computation to determine sign for comparison
immler
parents: 62420
diff changeset
  1794
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1795
end
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1796
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  1797
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1798
subsection \<open>Lemmas needed by Approximate\<close>
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1799
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1800
lemma Float_num[simp]:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1801
   "real_of_float (Float 1 0) = 1"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1802
   "real_of_float (Float 1 1) = 2"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1803
   "real_of_float (Float 1 2) = 4"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1804
   "real_of_float (Float 1 (- 1)) = 1/2"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1805
   "real_of_float (Float 1 (- 2)) = 1/4"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1806
   "real_of_float (Float 1 (- 3)) = 1/8"
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1807
   "real_of_float (Float (- 1) 0) = -1"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1808
   "real_of_float (Float (numeral n) 0) = numeral n"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1809
   "real_of_float (Float (- numeral n) 0) = - numeral n"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1810
  using two_powr_int_float[of 2] two_powr_int_float[of "-1"] two_powr_int_float[of "-2"]
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1811
    two_powr_int_float[of "-3"]
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1812
  using powr_realpow[of 2 2] powr_realpow[of 2 3]
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  1813
  using powr_minus[of "2::real" 1] powr_minus[of "2::real" 2] powr_minus[of "2::real" 3]
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1814
  by auto
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1815
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1816
lemma real_of_Float_int[simp]: "real_of_float (Float n 0) = real n"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1817
  by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1818
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1819
lemma float_zero[simp]: "real_of_float (Float 0 e) = 0"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1820
  by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1821
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1822
lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> \<bar>(a::int) div 2\<bar> < \<bar>a\<bar>"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1823
  by arith
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1824
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1825
lemma lapprox_rat: "real_of_float (lapprox_rat prec x y) \<le> real_of_int x / real_of_int y"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1826
  by (simp add: lapprox_rat.rep_eq truncate_down)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1827
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1828
lemma mult_div_le:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1829
  fixes a b :: int
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1830
  assumes "b > 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1831
  shows "a \<ge> b * (a div b)"
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1832
proof -
64246
15d1ee6e847b eliminated irregular aliasses
haftmann
parents: 64240
diff changeset
  1833
  from minus_div_mult_eq_mod [symmetric, of a b] have "a = b * (a div b) + a mod b"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1834
    by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1835
  also have "\<dots> \<ge> b * (a div b) + 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1836
    apply (rule add_left_mono)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1837
    apply (rule pos_mod_sign)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1838
    using assms
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1839
    apply simp
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1840
    done
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1841
  finally show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1842
    by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1843
qed
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1844
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1845
lemma lapprox_rat_nonneg:
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1846
  assumes "0 \<le> x" and "0 \<le> y"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1847
  shows "0 \<le> real_of_float (lapprox_rat n x y)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1848
  using assms
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1849
  by transfer (simp add: truncate_down_nonneg)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1850
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1851
lemma rapprox_rat: "real_of_int x / real_of_int y \<le> real_of_float (rapprox_rat prec x y)"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1852
  by transfer (simp add: truncate_up)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1853
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1854
lemma rapprox_rat_le1:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1855
  assumes "0 \<le> x" "0 < y" "x \<le> y"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1856
  shows "real_of_float (rapprox_rat n x y) \<le> 1"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1857
  using assms
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1858
  by transfer (simp add: truncate_up_le1)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1859
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1860
lemma rapprox_rat_nonneg_nonpos: "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real_of_float (rapprox_rat n x y) \<le> 0"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1861
  by transfer (simp add: truncate_up_nonpos divide_nonneg_nonpos)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1862
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1863
lemma rapprox_rat_nonpos_nonneg: "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real_of_float (rapprox_rat n x y) \<le> 0"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1864
  by transfer (simp add: truncate_up_nonpos divide_nonpos_nonneg)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1865
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1866
lemma real_divl: "real_divl prec x y \<le> x / y"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1867
  by (simp add: real_divl_def truncate_down)
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1868
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1869
lemma real_divr: "x / y \<le> real_divr prec x y"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1870
  by (simp add: real_divr_def truncate_up)
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1871
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1872
lemma float_divl: "real_of_float (float_divl prec x y) \<le> x / y"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1873
  by transfer (rule real_divl)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1874
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1875
lemma real_divl_lower_bound: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_divl prec x y"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1876
  by (simp add: real_divl_def truncate_down_nonneg)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1877
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1878
lemma float_divl_lower_bound: "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> real_of_float (float_divl prec x y)"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1879
  by transfer (rule real_divl_lower_bound)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1880
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1881
lemma exponent_1: "exponent 1 = 0"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1882
  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
  1883
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1884
lemma mantissa_1: "mantissa 1 = 1"
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1885
  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
  1886
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1887
lemma bitlen_1: "bitlen 1 = 1"
63248
414e3550e9c0 generalized bitlen to floor of log
immler
parents: 63040
diff changeset
  1888
  by (simp add: bitlen_alt_def)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1889
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1890
lemma float_upper_bound: "x \<le> 2 powr (bitlen \<bar>mantissa x\<bar> + exponent x)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1891
proof (cases "x = 0")
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1892
  case True
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1893
  then show ?thesis by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1894
next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1895
  case False
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1896
  then have "mantissa x \<noteq> 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1897
    using mantissa_eq_zero_iff by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1898
  have "x = mantissa x * 2 powr (exponent x)"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1899
    by (rule mantissa_exponent)
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1900
  also have "mantissa x \<le> \<bar>mantissa x\<bar>"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1901
    by simp
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1902
  also have "\<dots> \<le> 2 powr (bitlen \<bar>mantissa x\<bar>)"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1903
    using bitlen_bounds[of "\<bar>mantissa x\<bar>"] bitlen_nonneg \<open>mantissa x \<noteq> 0\<close>
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61639
diff changeset
  1904
    by (auto simp del: of_int_abs simp add: powr_int)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1905
  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
  1906
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1907
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1908
lemma real_divl_pos_less1_bound:
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1909
  assumes "0 < x" "x \<le> 1"
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1910
  shows "1 \<le> real_divl prec 1 x"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1911
  using assms
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1912
  by (auto intro!: truncate_down_ge1 simp: real_divl_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1913
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1914
lemma float_divl_pos_less1_bound:
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1915
  "0 < real_of_float x \<Longrightarrow> real_of_float x \<le> 1 \<Longrightarrow> prec \<ge> 1 \<Longrightarrow>
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1916
    1 \<le> real_of_float (float_divl prec 1 x)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1917
  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
  1918
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1919
lemma float_divr: "real_of_float x / real_of_float y \<le> real_of_float (float_divr prec x y)"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1920
  by transfer (rule real_divr)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1921
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1922
lemma real_divr_pos_less1_lower_bound:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1923
  assumes "0 < x"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1924
    and "x \<le> 1"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1925
  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
  1926
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1927
  have "1 \<le> 1 / x"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1928
    using \<open>0 < x\<close> and \<open>x \<le> 1\<close> by auto
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1929
  also have "\<dots> \<le> real_divr prec 1 x"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1930
    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
  1931
  finally show ?thesis by auto
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1932
qed
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  1933
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  1934
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
  1935
  by transfer (rule real_divr_pos_less1_lower_bound)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1936
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1937
lemma real_divr_nonpos_pos_upper_bound: "x \<le> 0 \<Longrightarrow> 0 \<le> y \<Longrightarrow> real_divr prec x y \<le> 0"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1938
  by (simp add: real_divr_def truncate_up_nonpos divide_le_0_iff)
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1939
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1940
lemma float_divr_nonpos_pos_upper_bound:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1941
  "real_of_float x \<le> 0 \<Longrightarrow> 0 \<le> real_of_float y \<Longrightarrow> real_of_float (float_divr prec x y) \<le> 0"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1942
  by transfer (rule real_divr_nonpos_pos_upper_bound)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1943
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1944
lemma real_divr_nonneg_neg_upper_bound: "0 \<le> x \<Longrightarrow> y \<le> 0 \<Longrightarrow> real_divr prec x y \<le> 0"
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1945
  by (simp add: real_divr_def truncate_up_nonpos divide_le_0_iff)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  1946
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  1947
lemma float_divr_nonneg_neg_upper_bound:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1948
  "0 \<le> real_of_float x \<Longrightarrow> real_of_float y \<le> 0 \<Longrightarrow> real_of_float (float_divr prec x y) \<le> 0"
54782
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1949
  by transfer (rule real_divr_nonneg_neg_upper_bound)
cd8f55c358c5 additional definitions and lemmas for Float
immler
parents: 54489
diff changeset
  1950
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1951
lemma truncate_up_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_up prec x \<le> truncate_up prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1954
proof -
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1955
  consider "\<lfloor>log 2 x\<rfloor> = \<lfloor>log 2 y\<rfloor>" | "\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>" "0 < x" | "x \<le> 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1956
    by arith
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1957
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1958
  proof cases
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1959
    case 1
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1960
    then show ?thesis
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1961
      using assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1962
      by (auto simp: truncate_up_def round_up_def intro!: ceiling_mono)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1963
  next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1964
    case 2
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1965
    from assms \<open>0 < x\<close> have "log 2 x \<le> log 2 y"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1966
      by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  1967
    with \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close>
70347
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
  1968
    have logless: "log 2 x < log 2 y"
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
  1969
      by linarith
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
  1970
    have flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 69593
diff changeset
  1971
      using \<open>\<lfloor>log 2 x\<rfloor> \<noteq> \<lfloor>log 2 y\<rfloor>\<close> \<open>log 2 x \<le> log 2 y\<close> by linarith
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1972
    have "truncate_up prec x =
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1973
      real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor> )\<rceil> * 2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1974
      using assms by (simp add: truncate_up_def round_up_def)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1975
    also have "\<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> (2 ^ (Suc prec))"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1976
    proof (simp only: ceiling_le_iff)
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1977
      have "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le>
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1978
        x * (2 powr real (Suc prec) / (2 powr log 2 x))"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1979
        using real_of_int_floor_add_one_ge[of "log 2 x"] assms
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  1980
        by (auto simp: algebra_simps simp flip: powr_diff intro!: mult_left_mono)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1981
      then show "x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> real_of_int ((2::int) ^ (Suc prec))"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1982
        using \<open>0 < x\<close> by (simp add: powr_realpow powr_add)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1983
    qed
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1984
    then have "real_of_int \<lceil>x * 2 powr real_of_int (int prec - \<lfloor>log 2 x\<rfloor>)\<rceil> \<le> 2 powr int (Suc prec)"
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1985
      by (auto simp: powr_realpow powr_add)
66912
a99a7cbf0fb5 generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
immler
parents: 65583
diff changeset
  1986
        (metis power_Suc of_int_le_numeral_power_cancel_iff)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1987
    also
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  1988
    have "2 powr - real_of_int (int prec - \<lfloor>log 2 x\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1989
      using logless flogless by (auto intro!: floor_mono)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1990
    also have "2 powr real_of_int (int (Suc prec)) \<le>
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1991
        2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1))"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  1992
      using assms \<open>0 < x\<close>
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1993
      by (auto simp: algebra_simps)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1994
    finally have "truncate_up prec x \<le>
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  1995
        2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 y\<rfloor> + 1)"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1996
      by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  1997
    also have "\<dots> = 2 powr (log 2 y + real_of_int (int prec - \<lfloor>log 2 y\<rfloor>) - real_of_int (int prec - \<lfloor>log 2 y\<rfloor>))"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1998
      by (subst powr_add[symmetric]) simp
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  1999
    also have "\<dots> = y"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2000
      using \<open>0 < x\<close> assms
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2001
      by (simp add: powr_add)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2002
    also have "\<dots> \<le> truncate_up prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2003
      by (rule truncate_up)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2004
    finally show ?thesis .
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2005
  next
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2006
    case 3
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2007
    then show ?thesis
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2008
      using assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2009
      by (auto intro!: truncate_up_le)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2010
  qed
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2011
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2012
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2013
lemma truncate_up_switch_sign_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2014
  assumes "x \<le> 0" "0 \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2015
  shows "truncate_up prec x \<le> truncate_up prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2016
proof -
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2017
  note truncate_up_nonpos[OF \<open>x \<le> 0\<close>]
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2018
  also note truncate_up_le[OF \<open>0 \<le> y\<close>]
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2019
  finally show ?thesis .
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2020
qed
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_down_switch_sign_mono:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2023
  assumes "x \<le> 0"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2024
    and "0 \<le> y"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2025
    and "x \<le> y"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2026
  shows "truncate_down prec x \<le> truncate_down prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2027
proof -
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2028
  note truncate_down_le[OF \<open>x \<le> 0\<close>]
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2029
  also note truncate_down_nonneg[OF \<open>0 \<le> y\<close>]
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2030
  finally show ?thesis .
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2031
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2032
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2033
lemma truncate_down_nonneg_mono:
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2034
  assumes "0 \<le> x" "x \<le> y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2035
  shows "truncate_down prec x \<le> truncate_down prec y"
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2036
proof -
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2037
  consider "x \<le> 0" | "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> = \<lfloor>log 2 \<bar>y\<bar>\<rfloor>" |
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2038
    "0 < x" "\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2039
    by arith
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2040
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2041
  proof cases
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2042
    case 1
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2043
    with assms have "x = 0" "0 \<le> y" by simp_all
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2044
    then show ?thesis
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2045
      by (auto intro!: truncate_down_nonneg)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2046
  next
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2047
    case 2
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2048
    then show ?thesis
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2049
      using assms
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2050
      by (auto simp: truncate_down_def round_down_def intro!: floor_mono)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2051
  next
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2052
    case 3
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2053
    from \<open>0 < x\<close> have "log 2 x \<le> log 2 y" "0 < y" "0 \<le> y"
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2054
      using assms by auto
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2055
    with \<open>\<lfloor>log 2 \<bar>x\<bar>\<rfloor> \<noteq> \<lfloor>log 2 \<bar>y\<bar>\<rfloor>\<close>
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2056
    have logless: "log 2 x < log 2 y" and flogless: "\<lfloor>log 2 x\<rfloor> < \<lfloor>log 2 y\<rfloor>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2057
      unfolding atomize_conj abs_of_pos[OF \<open>0 < x\<close>] abs_of_pos[OF \<open>0 < y\<close>]
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2058
      by (metis floor_less_cancel linorder_cases not_le)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2059
    have "2 powr prec \<le> y * 2 powr real prec / (2 powr log 2 y)"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2060
      using \<open>0 < y\<close> by simp
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2061
    also have "\<dots> \<le> y * 2 powr real (Suc prec) / (2 powr (real_of_int \<lfloor>log 2 y\<rfloor> + 1))"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2062
      using \<open>0 \<le> y\<close> \<open>0 \<le> x\<close> assms(2)
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  2063
      by (auto intro!: powr_mono divide_left_mono
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  2064
          simp: of_nat_diff powr_add powr_diff)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2065
    also have "\<dots> = y * 2 powr real (Suc prec) / (2 powr real_of_int \<lfloor>log 2 y\<rfloor> * 2)"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2066
      by (auto simp: powr_add)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2067
    finally have "(2 ^ prec) \<le> \<lfloor>y * 2 powr real_of_int (int (Suc prec) - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> - 1)\<rfloor>"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2068
      using \<open>0 \<le> y\<close>
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65109
diff changeset
  2069
      by (auto simp: powr_diff le_floor_iff powr_realpow powr_add)
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2070
    then have "(2 ^ (prec)) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>) \<le> truncate_down prec y"
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2071
      by (auto simp: truncate_down_def round_down_def)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  2072
    moreover have "x \<le> (2 ^ prec) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor>)"
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  2073
    proof -
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2074
      have "x = 2 powr (log 2 \<bar>x\<bar>)" using \<open>0 < x\<close> by simp
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2075
      also have "\<dots> \<le> (2 ^ (Suc prec )) * 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>)"
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2076
        using real_of_int_floor_add_one_ge[of "log 2 \<bar>x\<bar>"] \<open>0 < x\<close>
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  2077
        by (auto simp flip: powr_realpow powr_add simp: algebra_simps powr_mult_base le_powr_iff)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2078
      also
62420
c7666166c24e positive precision for truncate; fixed precision for approximation of rationals; code for truncate
immler
parents: 62419
diff changeset
  2079
      have "2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>x\<bar>\<rfloor>) \<le> 2 powr - real_of_int (int prec - \<lfloor>log 2 \<bar>y\<bar>\<rfloor> + 1)"
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60376
diff changeset
  2080
        using logless flogless \<open>x > 0\<close> \<open>y > 0\<close>
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2081
        by (auto intro!: floor_mono)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  2082
      finally show ?thesis
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  2083
        by (auto simp flip: powr_realpow simp: powr_diff assms of_nat_diff)
63356
77332fed33c3 misc tuning and modernization;
wenzelm
parents: 63248
diff changeset
  2084
    qed
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2085
    ultimately show ?thesis
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2086
      by (metis dual_order.trans truncate_down)
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2087
  qed
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2088
qed
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2089
58982
27e7e3f9e665 simplified computations based on round_up by reducing to round_down;
immler
parents: 58881
diff changeset
  2090
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
  2091
  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
  2092
  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
  2093
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2094
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
  2095
  apply (cases "0 \<le> x")
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2096
  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
  2097
  apply (simp add: truncate_down_eq_truncate_up)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2098
  apply (cases "0 \<le> y")
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2099
  apply (auto intro: truncate_up_nonneg_mono truncate_up_switch_sign_mono)
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2100
  done
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2101
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2102
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
  2103
  by (simp add: truncate_up_eq_truncate_down truncate_down_mono)
54784
54f1ce13c140 monotonicity of rounding and truncating float
immler
parents: 54783
diff changeset
  2104
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2105
lemma Float_le_zero_iff: "Float a b \<le> 0 \<longleftrightarrow> a \<le> 0"
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
  2106
  by (auto simp: zero_float_def mult_le_0_iff)
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2107
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2108
lemma real_of_float_pprt[simp]:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2109
  fixes a :: float
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2110
  shows "real_of_float (pprt a) = pprt (real_of_float a)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2111
  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
  2112
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2113
lemma real_of_float_nprt[simp]:
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2114
  fixes a :: float
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2115
  shows "real_of_float (nprt a) = nprt (real_of_float a)"
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2116
  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
  2117
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2118
context
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2119
begin
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2120
55565
f663fc1e653b simplify proofs because of the stronger reflexivity prover
kuncar
parents: 54784
diff changeset
  2121
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
  2122
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2123
qualified 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
  2124
  "int_floor_fl (Float m e) = (if 0 \<le> e then m * 2 ^ nat e else m div (2 ^ (nat (-e))))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2125
  apply transfer
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2126
  apply (simp add: powr_int floor_divide_of_int_eq)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
  2127
  apply (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power floor_of_int of_int_mult)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
  2128
  done
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2129
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
  2130
lift_definition floor_fl :: "float \<Rightarrow> float" is "\<lambda>x. real_of_int \<lfloor>x\<rfloor>"
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
  2131
  by simp
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2132
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2133
qualified 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
  2134
  "floor_fl (Float m e) = (if 0 \<le> e then Float m e else Float (m div (2 ^ (nat (-e)))) 0)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2135
  apply transfer
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2136
  apply (simp add: powr_int floor_divide_of_int_eq)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
  2137
  apply (metis (no_types, hide_lams)floor_divide_of_int_eq of_int_numeral of_int_power of_int_mult)
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61799
diff changeset
  2138
  done
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2139
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2140
end
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2141
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2142
lemma floor_fl: "real_of_float (floor_fl x) \<le> real_of_float x"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2143
  by transfer simp
47600
e12289b5796b use lifting to introduce floating point numbers
hoelzl
parents: 47599
diff changeset
  2144
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2145
lemma int_floor_fl: "real_of_int (int_floor_fl x) \<le> real_of_float x"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2146
  by transfer simp
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
  2147
47599
400b158f1589 replace the float datatype by a type with unique representation
hoelzl
parents: 47230
diff changeset
  2148
lemma floor_pos_exp: "exponent (floor_fl x) \<ge> 0"
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
  2149
proof (cases "floor_fl x = 0")
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2150
  case True
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2151
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2152
    by (simp add: floor_fl_def)
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2153
next
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2154
  case False
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2155
  have eq: "floor_fl x = Float \<lfloor>real_of_float x\<rfloor> 0"
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2156
    by transfer simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 60868
diff changeset
  2157
  obtain i where "\<lfloor>real_of_float x\<rfloor> = mantissa (floor_fl x) * 2 ^ i" "0 = exponent (floor_fl x) - int i"
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
  2158
    by (rule denormalize_shift[OF eq False])
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2159
  then show ?thesis
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2160
    by simp
53381
355a4cac5440 tuned proofs -- less guessing;
wenzelm
parents: 53215
diff changeset
  2161
qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2162
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2163
lemma compute_mantissa[code]:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2164
  "mantissa (Float m e) =
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2165
    (if m = 0 then 0 else if 2 dvd m then mantissa (normfloat (Float m e)) else m)"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  2166
  by (auto simp: mantissa_float Float.abs_eq simp flip: zero_float_def)
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2167
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2168
lemma compute_exponent[code]:
60698
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2169
  "exponent (Float m e) =
29e8bdc41f90 tuned proofs;
wenzelm
parents: 60679
diff changeset
  2170
    (if m = 0 then 0 else if 2 dvd m then exponent (normfloat (Float m e)) else e)"
68406
6beb45f6cf67 utilize 'flip'
nipkow
parents: 67573
diff changeset
  2171
  by (auto simp: exponent_float Float.abs_eq simp flip: zero_float_def)
67573
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
  2172
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
  2173
lifting_update Float.float.lifting
ed0a7090167d added lemmas, avoid 'float_of 0'
immler
parents: 67399
diff changeset
  2174
lifting_forget Float.float.lifting
58985
bf498e0af9e3 truncate intermediate results in horner to improve performance of approximate;
immler
parents: 58982
diff changeset
  2175
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
  2176
end