| author | wenzelm | 
| Sat, 03 Mar 2018 22:33:25 +0100 | |
| changeset 67763 | f4b1cf9e7010 | 
| parent 67573 | ed0a7090167d | 
| child 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1 | (* | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2 | Author: Johannes Hoelzl, TU Muenchen | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 3 | Coercions removed by Dmitriy Traytel | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 4 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 5 | This file contains only general material about computing lower/upper bounds | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 6 | on real functions. Approximation.thy contains the actual approximation algorithm | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 7 | and the approximation oracle. This is in order to make a clear separation between | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 8 | "morally immaculate" material about upper/lower bounds and the trusted oracle/reflection. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 9 | *) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 10 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 11 | theory Approximation_Bounds | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 12 | imports | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 13 | Complex_Main | 
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66280diff
changeset | 14 | "HOL-Library.Float" | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 15 | Dense_Linear_Order | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 16 | begin | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 17 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 18 | declare powr_neg_one [simp] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 19 | declare powr_neg_numeral [simp] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 20 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 21 | section "Horner Scheme" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 22 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 23 | subsection \<open>Define auxiliary helper \<open>horner\<close> function\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 24 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 25 | primrec horner :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat \<Rightarrow> nat) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> real" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 26 | "horner F G 0 i k x = 0" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 27 | "horner F G (Suc n) i k x = 1 / k - x * horner F G n (F i) (G i k) x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 28 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 29 | lemma horner_schema': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 30 | fixes x :: real and a :: "nat \<Rightarrow> real" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 31 | shows "a 0 - x * (\<Sum> i=0..<n. (-1)^i * a (Suc i) * x^i) = (\<Sum> i=0..<Suc n. (-1)^i * a i * x^i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 32 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 33 | have shift_pow: "\<And>i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 34 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 35 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 36 | unfolding sum_distrib_left shift_pow uminus_add_conv_diff [symmetric] sum_negf[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 37 | sum_head_upt_Suc[OF zero_less_Suc] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 38 | sum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n *a n * x^n"] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 39 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 40 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 41 | lemma horner_schema: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 42 | fixes f :: "nat \<Rightarrow> nat" and G :: "nat \<Rightarrow> nat \<Rightarrow> nat" and F :: "nat \<Rightarrow> nat" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 43 | assumes f_Suc: "\<And>n. f (Suc n) = G ((F ^^ n) s) (f n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 44 | shows "horner F G n ((F ^^ j') s) (f j') x = (\<Sum> j = 0..< n. (- 1) ^ j * (1 / (f (j' + j))) * x ^ j)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 45 | proof (induct n arbitrary: j') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 46 | case 0 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 47 | then show ?case by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 48 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 49 | case (Suc n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 50 | show ?case unfolding horner.simps Suc[where j'="Suc j'", unfolded funpow.simps comp_def f_Suc] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 51 | using horner_schema'[of "\<lambda> j. 1 / (f (j' + j))"] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 52 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 53 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 54 | lemma horner_bounds': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 55 | fixes lb :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" and ub :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 56 | assumes "0 \<le> real_of_float x" and f_Suc: "\<And>n. f (Suc n) = G ((F ^^ n) s) (f n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 57 | and lb_0: "\<And> i k x. lb 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 58 | and lb_Suc: "\<And> n i k x. lb (Suc n) i k x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 59 | (lapprox_rat prec 1 k) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 60 | (- float_round_up prec (x * (ub n (F i) (G i k) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 61 | and ub_0: "\<And> i k x. ub 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 62 | and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 63 | (rapprox_rat prec 1 k) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 64 | (- float_round_down prec (x * (lb n (F i) (G i k) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 65 | shows "(lb n ((F ^^ j') s) (f j') x) \<le> horner F G n ((F ^^ j') s) (f j') x \<and> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 66 | horner F G n ((F ^^ j') s) (f j') x \<le> (ub n ((F ^^ j') s) (f j') x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 67 | (is "?lb n j' \<le> ?horner n j' \<and> ?horner n j' \<le> ?ub n j'") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 68 | proof (induct n arbitrary: j') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 69 | case 0 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 70 | thus ?case unfolding lb_0 ub_0 horner.simps by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 71 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 72 | case (Suc n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 73 | thus ?case using lapprox_rat[of prec 1 "f j'"] using rapprox_rat[of 1 "f j'" prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 74 | Suc[where j'="Suc j'"] \<open>0 \<le> real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 75 | by (auto intro!: add_mono mult_left_mono float_round_down_le float_round_up_le | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 76 | order_trans[OF add_mono[OF _ float_plus_down_le]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 77 | order_trans[OF _ add_mono[OF _ float_plus_up_le]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 78 | simp add: lb_Suc ub_Suc field_simps f_Suc) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 79 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 80 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 81 | subsection "Theorems for floating point functions implementing the horner scheme" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 82 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 83 | text \<open> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 84 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 85 | Here @{term_type "f :: nat \<Rightarrow> nat"} is the sequence defining the Taylor series, the coefficients are
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 86 | all alternating and reciprocs. We use @{term G} and @{term F} to describe the computation of @{term f}.
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 87 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 88 | \<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 89 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 90 | lemma horner_bounds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 91 | fixes F :: "nat \<Rightarrow> nat" and G :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 92 | assumes "0 \<le> real_of_float x" and f_Suc: "\<And>n. f (Suc n) = G ((F ^^ n) s) (f n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 93 | and lb_0: "\<And> i k x. lb 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 94 | and lb_Suc: "\<And> n i k x. lb (Suc n) i k x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 95 | (lapprox_rat prec 1 k) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 96 | (- float_round_up prec (x * (ub n (F i) (G i k) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 97 | and ub_0: "\<And> i k x. ub 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 98 | and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 99 | (rapprox_rat prec 1 k) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 100 | (- float_round_down prec (x * (lb n (F i) (G i k) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 101 | shows "(lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (- 1) ^ j * (1 / (f (j' + j))) * (x ^ j))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 102 | (is "?lb") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 103 | and "(\<Sum>j=0..<n. (- 1) ^ j * (1 / (f (j' + j))) * (x ^ j)) \<le> (ub n ((F ^^ j') s) (f j') x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 104 | (is "?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 105 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 106 | have "?lb \<and> ?ub" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 107 | using horner_bounds'[where lb=lb, OF \<open>0 \<le> real_of_float x\<close> f_Suc lb_0 lb_Suc ub_0 ub_Suc] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 108 | unfolding horner_schema[where f=f, OF f_Suc] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 109 | thus "?lb" and "?ub" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 110 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 111 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 112 | lemma horner_bounds_nonpos: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 113 | fixes F :: "nat \<Rightarrow> nat" and G :: "nat \<Rightarrow> nat \<Rightarrow> nat" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 114 | assumes "real_of_float x \<le> 0" and f_Suc: "\<And>n. f (Suc n) = G ((F ^^ n) s) (f n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 115 | and lb_0: "\<And> i k x. lb 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 116 | and lb_Suc: "\<And> n i k x. lb (Suc n) i k x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 117 | (lapprox_rat prec 1 k) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 118 | (float_round_down prec (x * (ub n (F i) (G i k) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 119 | and ub_0: "\<And> i k x. ub 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 120 | and ub_Suc: "\<And> n i k x. ub (Suc n) i k x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 121 | (rapprox_rat prec 1 k) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 122 | (float_round_up prec (x * (lb n (F i) (G i k) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 123 | shows "(lb n ((F ^^ j') s) (f j') x) \<le> (\<Sum>j=0..<n. (1 / (f (j' + j))) * real_of_float x ^ j)" (is "?lb") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 124 | and "(\<Sum>j=0..<n. (1 / (f (j' + j))) * real_of_float x ^ j) \<le> (ub n ((F ^^ j') s) (f j') x)" (is "?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 125 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 126 | have diff_mult_minus: "x - y * z = x + - y * z" for x y z :: float by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 127 | have sum_eq: "(\<Sum>j=0..<n. (1 / (f (j' + j))) * real_of_float x ^ j) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 128 | (\<Sum>j = 0..<n. (- 1) ^ j * (1 / (f (j' + j))) * real_of_float (- x) ^ j)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 129 | by (auto simp add: field_simps power_mult_distrib[symmetric]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 130 | have "0 \<le> real_of_float (-x)" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 131 | from horner_bounds[where G=G and F=F and f=f and s=s and prec=prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 132 | and lb="\<lambda> n i k x. lb n i k (-x)" and ub="\<lambda> n i k x. ub n i k (-x)", | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 133 | unfolded lb_Suc ub_Suc diff_mult_minus, | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 134 | OF this f_Suc lb_0 _ ub_0 _] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 135 | show "?lb" and "?ub" unfolding minus_minus sum_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 136 | by (auto simp: minus_float_round_up_eq minus_float_round_down_eq) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 137 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 138 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 139 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 140 | subsection \<open>Selectors for next even or odd number\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 141 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 142 | text \<open> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 143 | The horner scheme computes alternating series. To get the upper and lower bounds we need to | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 144 | guarantee to access a even or odd member. To do this we use @{term get_odd} and @{term get_even}.
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 145 | \<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 146 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 147 | definition get_odd :: "nat \<Rightarrow> nat" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 148 | "get_odd n = (if odd n then n else (Suc n))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 149 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 150 | definition get_even :: "nat \<Rightarrow> nat" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 151 | "get_even n = (if even n then n else (Suc n))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 152 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 153 | lemma get_odd[simp]: "odd (get_odd n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 154 | unfolding get_odd_def by (cases "odd n") auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 155 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 156 | lemma get_even[simp]: "even (get_even n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 157 | unfolding get_even_def by (cases "even n") auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 158 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 159 | lemma get_odd_ex: "\<exists> k. Suc k = get_odd n \<and> odd (Suc k)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 160 | by (auto simp: get_odd_def odd_pos intro!: exI[of _ "n - 1"]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 161 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 162 | lemma get_even_double: "\<exists>i. get_even n = 2 * i" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 163 | using get_even by (blast elim: evenE) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 164 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 165 | lemma get_odd_double: "\<exists>i. get_odd n = 2 * i + 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 166 | using get_odd by (blast elim: oddE) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 167 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 168 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 169 | section "Power function" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 170 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 171 | definition float_power_bnds :: "nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float * float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 172 | "float_power_bnds prec n l u = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 173 | (if 0 < l then (power_down_fl prec l n, power_up_fl prec u n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 174 | else if odd n then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 175 | (- power_up_fl prec \<bar>l\<bar> n, | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 176 | if u < 0 then - power_down_fl prec \<bar>u\<bar> n else power_up_fl prec u n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 177 | else if u < 0 then (power_down_fl prec \<bar>u\<bar> n, power_up_fl prec \<bar>l\<bar> n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 178 | else (0, power_up_fl prec (max \<bar>l\<bar> \<bar>u\<bar>) n))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 179 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 180 | lemma le_minus_power_downI: "0 \<le> x \<Longrightarrow> x ^ n \<le> - a \<Longrightarrow> a \<le> - power_down prec x n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 181 | by (subst le_minus_iff) (auto intro: power_down_le power_mono_odd) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 182 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 183 | lemma float_power_bnds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 184 |   "(l1, u1) = float_power_bnds prec n l u \<Longrightarrow> x \<in> {l .. u} \<Longrightarrow> (x::real) ^ n \<in> {l1..u1}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 185 | by (auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 186 | simp: float_power_bnds_def max_def real_power_up_fl real_power_down_fl minus_le_iff | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 187 | split: if_split_asm | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 188 | intro!: power_up_le power_down_le le_minus_power_downI | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 189 | intro: power_mono_odd power_mono power_mono_even zero_le_even_power) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 190 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 191 | lemma bnds_power: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 192 |   "\<forall>(x::real) l u. (l1, u1) = float_power_bnds prec n l u \<and> x \<in> {l .. u} \<longrightarrow>
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 193 | l1 \<le> x ^ n \<and> x ^ n \<le> u1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 194 | using float_power_bnds by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 195 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 196 | section \<open>Approximation utility functions\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 197 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 198 | definition bnds_mult :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<times> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 199 | "bnds_mult prec a1 a2 b1 b2 = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 200 | (float_plus_down prec (nprt a1 * pprt b2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 201 | (float_plus_down prec (nprt a2 * nprt b2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 202 | (float_plus_down prec (pprt a1 * pprt b1) (pprt a2 * nprt b1))), | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 203 | float_plus_up prec (pprt a2 * pprt b2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 204 | (float_plus_up prec (pprt a1 * nprt b2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 205 | (float_plus_up prec (nprt a2 * pprt b1) (nprt a1 * nprt b1))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 206 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 207 | lemma bnds_mult: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 208 | fixes prec :: nat and a1 aa2 b1 b2 :: float | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 209 | assumes "(l, u) = bnds_mult prec a1 a2 b1 b2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 210 |   assumes "a \<in> {real_of_float a1..real_of_float a2}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 211 |   assumes "b \<in> {real_of_float b1..real_of_float b2}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 212 |   shows   "a * b \<in> {real_of_float l..real_of_float u}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 213 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 214 | from assms have "real_of_float l \<le> a * b" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 215 | by (intro order.trans[OF _ mult_ge_prts[of a1 a a2 b1 b b2]]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 216 | (auto simp: bnds_mult_def intro!: float_plus_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 217 | moreover from assms have "real_of_float u \<ge> a * b" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 218 | by (intro order.trans[OF mult_le_prts[of a1 a a2 b1 b b2]]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 219 | (auto simp: bnds_mult_def intro!: float_plus_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 220 | ultimately show ?thesis by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 221 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 222 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 223 | definition map_bnds :: "(nat \<Rightarrow> float \<Rightarrow> float) \<Rightarrow> (nat \<Rightarrow> float \<Rightarrow> float) \<Rightarrow> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 224 | nat \<Rightarrow> (float \<times> float) \<Rightarrow> (float \<times> float)" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 225 | "map_bnds lb ub prec = (\<lambda>(l,u). (lb prec l, ub prec u))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 226 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 227 | lemma map_bnds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 228 | assumes "(lf, uf) = map_bnds lb ub prec (l, u)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 229 | assumes "mono f" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 230 |   assumes "x \<in> {real_of_float l..real_of_float u}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 231 | assumes "real_of_float (lb prec l) \<le> f (real_of_float l)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 232 | assumes "real_of_float (ub prec u) \<ge> f (real_of_float u)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 233 |   shows   "f x \<in> {real_of_float lf..real_of_float uf}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 234 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 235 | from assms have "real_of_float lf = real_of_float (lb prec l)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 236 | by (simp add: map_bnds_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 237 | also have "real_of_float (lb prec l) \<le> f (real_of_float l)" by fact | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 238 | also from assms have "\<dots> \<le> f x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 239 | by (intro monoD[OF \<open>mono f\<close>]) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 240 | finally have lf: "real_of_float lf \<le> f x" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 241 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 242 | from assms have "f x \<le> f (real_of_float u)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 243 | by (intro monoD[OF \<open>mono f\<close>]) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 244 | also have "\<dots> \<le> real_of_float (ub prec u)" by fact | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 245 | also from assms have "\<dots> = real_of_float uf" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 246 | by (simp add: map_bnds_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 247 | finally have uf: "f x \<le> real_of_float uf" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 248 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 249 | from lf uf show ?thesis by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 250 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 251 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 252 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 253 | section "Square root" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 254 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 255 | text \<open> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 256 | The square root computation is implemented as newton iteration. As first first step we use the | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 257 | nearest power of two greater than the square root. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 258 | \<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 259 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 260 | fun sqrt_iteration :: "nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 261 | "sqrt_iteration prec 0 x = Float 1 ((bitlen \<bar>mantissa x\<bar> + exponent x) div 2 + 1)" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 262 | "sqrt_iteration prec (Suc m) x = (let y = sqrt_iteration prec m x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 263 | in Float 1 (- 1) * float_plus_up prec y (float_divr prec x y))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 264 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 265 | lemma compute_sqrt_iteration_base[code]: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 266 | shows "sqrt_iteration prec n (Float m e) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 267 | (if n = 0 then Float 1 ((if m = 0 then 0 else bitlen \<bar>m\<bar> + e) div 2 + 1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 268 | else (let y = sqrt_iteration prec (n - 1) (Float m e) in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 269 | Float 1 (- 1) * float_plus_up prec y (float_divr prec (Float m e) y)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 270 | using bitlen_Float by (cases n) simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 271 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 272 | function ub_sqrt lb_sqrt :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 273 | "ub_sqrt prec x = (if 0 < x then (sqrt_iteration prec prec x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 274 | else if x < 0 then - lb_sqrt prec (- x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 275 | else 0)" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 276 | "lb_sqrt prec x = (if 0 < x then (float_divl prec x (sqrt_iteration prec prec x)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 277 | else if x < 0 then - ub_sqrt prec (- x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 278 | else 0)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 279 | by pat_completeness auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 280 | termination by (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 281 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 282 | declare lb_sqrt.simps[simp del] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 283 | declare ub_sqrt.simps[simp del] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 284 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 285 | lemma sqrt_ub_pos_pos_1: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 286 | assumes "sqrt x < b" and "0 < b" and "0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 287 | shows "sqrt x < (b + x / b)/2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 288 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 289 | from assms have "0 < (b - sqrt x)\<^sup>2 " by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 290 | also have "\<dots> = b\<^sup>2 - 2 * b * sqrt x + (sqrt x)\<^sup>2" by algebra | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 291 | also have "\<dots> = b\<^sup>2 - 2 * b * sqrt x + x" using assms by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 292 | finally have "0 < b\<^sup>2 - 2 * b * sqrt x + x" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 293 | hence "0 < b / 2 - sqrt x + x / (2 * b)" using assms | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 294 | by (simp add: field_simps power2_eq_square) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 295 | thus ?thesis by (simp add: field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 296 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 297 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 298 | lemma sqrt_iteration_bound: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 299 | assumes "0 < real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 300 | shows "sqrt x < sqrt_iteration prec n x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 301 | proof (induct n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 302 | case 0 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 303 | show ?case | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 304 | proof (cases x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 305 | case (Float m e) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 306 | hence "0 < m" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 307 | using assms | 
| 67573 | 308 | by (auto simp: sign_simps) | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 309 | hence "0 < sqrt m" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 310 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 311 | have int_nat_bl: "(nat (bitlen m)) = bitlen m" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 312 | using bitlen_nonneg by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 313 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 314 | have "x = (m / 2^nat (bitlen m)) * 2 powr (e + (nat (bitlen m)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 315 | unfolding Float by (auto simp: powr_realpow[symmetric] field_simps powr_add) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 316 | also have "\<dots> < 1 * 2 powr (e + nat (bitlen m))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 317 | proof (rule mult_strict_right_mono, auto) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 318 | show "m < 2^nat (bitlen m)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 319 | using bitlen_bounds[OF \<open>0 < m\<close>, THEN conjunct2] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 320 | unfolding of_int_less_iff[of m, symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 321 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 322 | finally have "sqrt x < sqrt (2 powr (e + bitlen m))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 323 | unfolding int_nat_bl by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 324 | also have "\<dots> \<le> 2 powr ((e + bitlen m) div 2 + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 325 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 326 | let ?E = "e + bitlen m" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 327 | have E_mod_pow: "2 powr (?E mod 2) < 4" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 328 | proof (cases "?E mod 2 = 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 329 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 330 | thus ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 331 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 332 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 333 | have "0 \<le> ?E mod 2" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 334 | have "?E mod 2 < 2" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 335 | from this[THEN zless_imp_add1_zle] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 336 | have "?E mod 2 \<le> 0" using False by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 337 | from xt1(5)[OF \<open>0 \<le> ?E mod 2\<close> this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 338 | show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 339 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 340 | hence "sqrt (2 powr (?E mod 2)) < sqrt (2 * 2)" | 
| 66280 
0c5eb47e2696
Adapted Approximation_Bounds to changes in Multiset
 eberlm <eberlm@in.tum.de> parents: 
65582diff
changeset | 341 | by (intro real_sqrt_less_mono) auto | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 342 | hence E_mod_pow: "sqrt (2 powr (?E mod 2)) < 2" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 343 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 344 | have E_eq: "2 powr ?E = 2 powr (?E div 2 + ?E div 2 + ?E mod 2)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 345 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 346 | have "sqrt (2 powr ?E) = sqrt (2 powr (?E div 2) * 2 powr (?E div 2) * 2 powr (?E mod 2))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 347 | unfolding E_eq unfolding powr_add[symmetric] by (metis of_int_add) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 348 | also have "\<dots> = 2 powr (?E div 2) * sqrt (2 powr (?E mod 2))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 349 | unfolding real_sqrt_mult[of _ "2 powr (?E mod 2)"] real_sqrt_abs2 by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 350 | also have "\<dots> < 2 powr (?E div 2) * 2 powr 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 351 | by (rule mult_strict_left_mono) (auto intro: E_mod_pow) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 352 | also have "\<dots> = 2 powr (?E div 2 + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 353 | unfolding add.commute[of _ 1] powr_add[symmetric] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 354 | finally show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 355 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 356 | finally show ?thesis using \<open>0 < m\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 357 | unfolding Float | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 358 | by (subst compute_sqrt_iteration_base) (simp add: ac_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 359 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 360 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 361 | case (Suc n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 362 | let ?b = "sqrt_iteration prec n x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 363 | have "0 < sqrt x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 364 | using \<open>0 < real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 365 | also have "\<dots> < real_of_float ?b" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 366 | using Suc . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 367 | finally have "sqrt x < (?b + x / ?b)/2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 368 | using sqrt_ub_pos_pos_1[OF Suc _ \<open>0 < real_of_float x\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 369 | also have "\<dots> \<le> (?b + (float_divr prec x ?b))/2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 370 | by (rule divide_right_mono, auto simp add: float_divr) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 371 | also have "\<dots> = (Float 1 (- 1)) * (?b + (float_divr prec x ?b))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 372 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 373 | also have "\<dots> \<le> (Float 1 (- 1)) * (float_plus_up prec ?b (float_divr prec x ?b))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 374 | by (auto simp add: algebra_simps float_plus_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 375 | finally show ?case | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 376 | unfolding sqrt_iteration.simps Let_def distrib_left . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 377 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 378 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 379 | lemma sqrt_iteration_lower_bound: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 380 | assumes "0 < real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 381 | shows "0 < real_of_float (sqrt_iteration prec n x)" (is "0 < ?sqrt") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 382 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 383 | have "0 < sqrt x" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 384 | also have "\<dots> < ?sqrt" using sqrt_iteration_bound[OF assms] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 385 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 386 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 387 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 388 | lemma lb_sqrt_lower_bound: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 389 | assumes "0 \<le> real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 390 | shows "0 \<le> real_of_float (lb_sqrt prec x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 391 | proof (cases "0 < x") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 392 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 393 | hence "0 < real_of_float x" and "0 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 394 | using \<open>0 \<le> real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 395 | hence "0 < sqrt_iteration prec prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 396 | using sqrt_iteration_lower_bound by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 397 | hence "0 \<le> real_of_float (float_divl prec x (sqrt_iteration prec prec x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 398 | using float_divl_lower_bound[OF \<open>0 \<le> x\<close>] unfolding less_eq_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 399 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 400 | unfolding lb_sqrt.simps using True by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 401 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 402 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 403 | with \<open>0 \<le> real_of_float x\<close> have "real_of_float x = 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 404 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 405 | unfolding lb_sqrt.simps by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 406 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 407 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 408 | lemma bnds_sqrt': "sqrt x \<in> {(lb_sqrt prec x) .. (ub_sqrt prec x)}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 409 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 410 | have lb: "lb_sqrt prec x \<le> sqrt x" if "0 < x" for x :: float | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 411 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 412 | from that have "0 < real_of_float x" and "0 \<le> real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 413 | hence sqrt_gt0: "0 < sqrt x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 414 | hence sqrt_ub: "sqrt x < sqrt_iteration prec prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 415 | using sqrt_iteration_bound by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 416 | have "(float_divl prec x (sqrt_iteration prec prec x)) \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 417 | x / (sqrt_iteration prec prec x)" by (rule float_divl) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 418 | also have "\<dots> < x / sqrt x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 419 | by (rule divide_strict_left_mono[OF sqrt_ub \<open>0 < real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 420 | mult_pos_pos[OF order_less_trans[OF sqrt_gt0 sqrt_ub] sqrt_gt0]]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 421 | also have "\<dots> = sqrt x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 422 | unfolding inverse_eq_iff_eq[of _ "sqrt x", symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 423 | sqrt_divide_self_eq[OF \<open>0 \<le> real_of_float x\<close>, symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 424 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 425 | unfolding lb_sqrt.simps if_P[OF \<open>0 < x\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 426 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 427 | have ub: "sqrt x \<le> ub_sqrt prec x" if "0 < x" for x :: float | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 428 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 429 | from that have "0 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 430 | hence "0 < sqrt x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 431 | hence "sqrt x < sqrt_iteration prec prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 432 | using sqrt_iteration_bound by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 433 | then show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 434 | unfolding ub_sqrt.simps if_P[OF \<open>0 < x\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 435 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 436 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 437 | using lb[of "-x"] ub[of "-x"] lb[of x] ub[of x] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 438 | by (auto simp add: lb_sqrt.simps ub_sqrt.simps real_sqrt_minus) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 439 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 440 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 441 | lemma bnds_sqrt: "\<forall>(x::real) lx ux. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 442 |   (l, u) = (lb_sqrt prec lx, ub_sqrt prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> sqrt x \<and> sqrt x \<le> u"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 443 | proof ((rule allI) +, rule impI, erule conjE, rule conjI) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 444 | fix x :: real | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 445 | fix lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 446 | assume "(l, u) = (lb_sqrt prec lx, ub_sqrt prec ux)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 447 |     and x: "x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 448 | hence l: "l = lb_sqrt prec lx " and u: "u = ub_sqrt prec ux" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 449 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 450 | have "sqrt lx \<le> sqrt x" using x by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 451 | from order_trans[OF _ this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 452 | show "l \<le> sqrt x" unfolding l using bnds_sqrt'[of lx prec] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 453 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 454 | have "sqrt x \<le> sqrt ux" using x by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 455 | from order_trans[OF this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 456 | show "sqrt x \<le> u" unfolding u using bnds_sqrt'[of ux prec] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 457 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 458 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 459 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 460 | section "Arcus tangens and \<pi>" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 461 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 462 | subsection "Compute arcus tangens series" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 463 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 464 | text \<open> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 465 | As first step we implement the computation of the arcus tangens series. This is only valid in the range | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 466 | @{term "{-1 :: real .. 1}"}. This is used to compute \<pi> and then the entire arcus tangens.
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 467 | \<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 468 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 469 | fun ub_arctan_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 470 | and lb_arctan_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 471 | "ub_arctan_horner prec 0 k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 472 | | "ub_arctan_horner prec (Suc n) k x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 473 | (rapprox_rat prec 1 k) (- float_round_down prec (x * (lb_arctan_horner prec n (k + 2) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 474 | | "lb_arctan_horner prec 0 k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 475 | | "lb_arctan_horner prec (Suc n) k x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 476 | (lapprox_rat prec 1 k) (- float_round_up prec (x * (ub_arctan_horner prec n (k + 2) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 477 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 478 | lemma arctan_0_1_bounds': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 479 | assumes "0 \<le> real_of_float y" "real_of_float y \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 480 | and "even n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 481 | shows "arctan (sqrt y) \<in> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 482 |       {(sqrt y * lb_arctan_horner prec n 1 y) .. (sqrt y * ub_arctan_horner prec (Suc n) 1 y)}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 483 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 484 | let ?c = "\<lambda>i. (- 1) ^ i * (1 / (i * 2 + (1::nat)) * sqrt y ^ (i * 2 + 1))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 485 | let ?S = "\<lambda>n. \<Sum> i=0..<n. ?c i" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 486 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 487 | have "0 \<le> sqrt y" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 488 | have "sqrt y \<le> 1" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 489 | from \<open>even n\<close> obtain m where "2 * m = n" by (blast elim: evenE) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 490 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 491 |   have "arctan (sqrt y) \<in> { ?S n .. ?S (Suc n) }"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 492 | proof (cases "sqrt y = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 493 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 494 | then show ?thesis by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 495 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 496 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 497 | hence "0 < sqrt y" using \<open>0 \<le> sqrt y\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 498 | hence prem: "0 < 1 / (0 * 2 + (1::nat)) * sqrt y ^ (0 * 2 + 1)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 499 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 500 | have "\<bar> sqrt y \<bar> \<le> 1" using \<open>0 \<le> sqrt y\<close> \<open>sqrt y \<le> 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 501 | from mp[OF summable_Leibniz(2)[OF zeroseq_arctan_series[OF this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 502 | monoseq_arctan_series[OF this]] prem, THEN spec, of m, unfolded \<open>2 * m = n\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 503 | show ?thesis unfolding arctan_series[OF \<open>\<bar> sqrt y \<bar> \<le> 1\<close>] Suc_eq_plus1 atLeast0LessThan . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 504 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 505 | note arctan_bounds = this[unfolded atLeastAtMost_iff] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 506 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 507 | have F: "\<And>n. 2 * Suc n + 1 = 2 * n + 1 + 2" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 508 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 509 | note bounds = horner_bounds[where s=1 and f="\<lambda>i. 2 * i + 1" and j'=0 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 510 | and lb="\<lambda>n i k x. lb_arctan_horner prec n k x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 511 | and ub="\<lambda>n i k x. ub_arctan_horner prec n k x", | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 512 | OF \<open>0 \<le> real_of_float y\<close> F lb_arctan_horner.simps ub_arctan_horner.simps] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 513 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 514 | have "(sqrt y * lb_arctan_horner prec n 1 y) \<le> arctan (sqrt y)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 515 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 516 | have "(sqrt y * lb_arctan_horner prec n 1 y) \<le> ?S n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 517 | using bounds(1) \<open>0 \<le> sqrt y\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 518 | apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 519 | apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 520 | apply (auto intro!: mult_left_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 521 | done | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 522 | also have "\<dots> \<le> arctan (sqrt y)" using arctan_bounds .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 523 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 524 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 525 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 526 | have "arctan (sqrt y) \<le> (sqrt y * ub_arctan_horner prec (Suc n) 1 y)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 527 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 528 | have "arctan (sqrt y) \<le> ?S (Suc n)" using arctan_bounds .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 529 | also have "\<dots> \<le> (sqrt y * ub_arctan_horner prec (Suc n) 1 y)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 530 | using bounds(2)[of "Suc n"] \<open>0 \<le> sqrt y\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 531 | apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 532 | apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 533 | apply (auto intro!: mult_left_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 534 | done | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 535 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 536 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 537 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 538 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 539 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 540 | lemma arctan_0_1_bounds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 541 | assumes "0 \<le> real_of_float y" "real_of_float y \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 542 | shows "arctan (sqrt y) \<in> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 543 |     {(sqrt y * lb_arctan_horner prec (get_even n) 1 y) ..
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 544 | (sqrt y * ub_arctan_horner prec (get_odd n) 1 y)}" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 545 | using | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 546 | arctan_0_1_bounds'[OF assms, of n prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 547 | arctan_0_1_bounds'[OF assms, of "n + 1" prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 548 | arctan_0_1_bounds'[OF assms, of "n - 1" prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 549 | by (auto simp: get_even_def get_odd_def odd_pos | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 550 | simp del: ub_arctan_horner.simps lb_arctan_horner.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 551 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 552 | lemma arctan_lower_bound: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 553 | assumes "0 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 554 | shows "x / (1 + x\<^sup>2) \<le> arctan x" (is "?l x \<le> _") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 555 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 556 | have "?l x - arctan x \<le> ?l 0 - arctan 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 557 | using assms | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 558 | by (intro DERIV_nonpos_imp_nonincreasing[where f="\<lambda>x. ?l x - arctan x"]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 559 | (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 560 | thus ?thesis by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 561 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 562 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 563 | lemma arctan_divide_mono: "0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> arctan y / y \<le> arctan x / x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 564 | by (rule DERIV_nonpos_imp_nonincreasing[where f="\<lambda>x. arctan x / x"]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 565 | (auto intro!: derivative_eq_intros divide_nonpos_nonneg | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 566 | simp: inverse_eq_divide arctan_lower_bound) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 567 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 568 | lemma arctan_mult_mono: "0 \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> x * arctan y \<le> y * arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 569 | using arctan_divide_mono[of x y] by (cases "x = 0") (simp_all add: field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 570 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 571 | lemma arctan_mult_le: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 572 | assumes "0 \<le> x" "x \<le> y" "y * z \<le> arctan y" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 573 | shows "x * z \<le> arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 574 | proof (cases "x = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 575 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 576 | then show ?thesis by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 577 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 578 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 579 | with assms have "z \<le> arctan y / y" by (simp add: field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 580 | also have "\<dots> \<le> arctan x / x" using assms \<open>x \<noteq> 0\<close> by (auto intro!: arctan_divide_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 581 | finally show ?thesis using assms \<open>x \<noteq> 0\<close> by (simp add: field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 582 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 583 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 584 | lemma arctan_le_mult: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 585 | assumes "0 < x" "x \<le> y" "arctan x \<le> x * z" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 586 | shows "arctan y \<le> y * z" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 587 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 588 | from assms have "arctan y / y \<le> arctan x / x" by (auto intro!: arctan_divide_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 589 | also have "\<dots> \<le> z" using assms by (auto simp: field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 590 | finally show ?thesis using assms by (simp add: field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 591 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 592 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 593 | lemma arctan_0_1_bounds_le: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 594 | assumes "0 \<le> x" "x \<le> 1" "0 < real_of_float xl" "real_of_float xl \<le> x * x" "x * x \<le> real_of_float xu" "real_of_float xu \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 595 | shows "arctan x \<in> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 596 |       {x * lb_arctan_horner p1 (get_even n) 1 xu .. x * ub_arctan_horner p2 (get_odd n) 1 xl}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 597 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 598 | from assms have "real_of_float xl \<le> 1" "sqrt (real_of_float xl) \<le> x" "x \<le> sqrt (real_of_float xu)" "0 \<le> real_of_float xu" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 599 | "0 \<le> real_of_float xl" "0 < sqrt (real_of_float xl)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 600 | by (auto intro!: real_le_rsqrt real_le_lsqrt simp: power2_eq_square) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 601 | from arctan_0_1_bounds[OF \<open>0 \<le> real_of_float xu\<close> \<open>real_of_float xu \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 602 | have "sqrt (real_of_float xu) * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \<le> arctan (sqrt (real_of_float xu))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 603 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 604 | from arctan_mult_le[OF \<open>0 \<le> x\<close> \<open>x \<le> sqrt _\<close> this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 605 | have "x * real_of_float (lb_arctan_horner p1 (get_even n) 1 xu) \<le> arctan x" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 606 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 607 | from arctan_0_1_bounds[OF \<open>0 \<le> real_of_float xl\<close> \<open>real_of_float xl \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 608 | have "arctan (sqrt (real_of_float xl)) \<le> sqrt (real_of_float xl) * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 609 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 610 | from arctan_le_mult[OF \<open>0 < sqrt xl\<close> \<open>sqrt xl \<le> x\<close> this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 611 | have "arctan x \<le> x * real_of_float (ub_arctan_horner p2 (get_odd n) 1 xl)" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 612 | ultimately show ?thesis by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 613 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 614 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 615 | lemma arctan_0_1_bounds_round: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 616 | assumes "0 \<le> real_of_float x" "real_of_float x \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 617 | shows "arctan x \<in> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 618 |       {real_of_float x * lb_arctan_horner p1 (get_even n) 1 (float_round_up (Suc p2) (x * x)) ..
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 619 | real_of_float x * ub_arctan_horner p3 (get_odd n) 1 (float_round_down (Suc p4) (x * x))}" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 620 | using assms | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 621 | apply (cases "x > 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 622 | apply (intro arctan_0_1_bounds_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 623 | apply (auto simp: float_round_down.rep_eq float_round_up.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 624 | intro!: truncate_up_le1 mult_le_one truncate_down_le truncate_up_le truncate_down_pos | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 625 | mult_pos_pos) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 626 | done | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 627 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 628 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 629 | subsection "Compute \<pi>" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 630 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 631 | definition ub_pi :: "nat \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 632 | "ub_pi prec = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 633 | (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 634 | A = rapprox_rat prec 1 5 ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 635 | B = lapprox_rat prec 1 239 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 636 | in ((Float 1 2) * float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 637 | ((Float 1 2) * float_round_up prec (A * (ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 638 | (float_round_down (Suc prec) (A * A))))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 639 | (- float_round_down prec (B * (lb_arctan_horner prec (get_even (prec div 14 + 1)) 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 640 | (float_round_up (Suc prec) (B * B)))))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 641 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 642 | definition lb_pi :: "nat \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 643 | "lb_pi prec = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 644 | (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 645 | A = lapprox_rat prec 1 5 ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 646 | B = rapprox_rat prec 1 239 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 647 | in ((Float 1 2) * float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 648 | ((Float 1 2) * float_round_down prec (A * (lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 649 | (float_round_up (Suc prec) (A * A))))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 650 | (- float_round_up prec (B * (ub_arctan_horner prec (get_odd (prec div 14 + 1)) 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 651 | (float_round_down (Suc prec) (B * B)))))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 652 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 653 | lemma pi_boundaries: "pi \<in> {(lb_pi n) .. (ub_pi n)}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 654 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 655 | have machin_pi: "pi = 4 * (4 * arctan (1 / 5) - arctan (1 / 239))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 656 | unfolding machin[symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 657 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 658 |   {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 659 | fix prec n :: nat | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 660 | fix k :: int | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 661 | assume "1 < k" hence "0 \<le> k" and "0 < k" and "1 \<le> k" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 662 | let ?k = "rapprox_rat prec 1 k" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 663 | let ?kl = "float_round_down (Suc prec) (?k * ?k)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 664 | have "1 div k = 0" using div_pos_pos_trivial[OF _ \<open>1 < k\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 665 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 666 | have "0 \<le> real_of_float ?k" by (rule order_trans[OF _ rapprox_rat]) (auto simp add: \<open>0 \<le> k\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 667 | have "real_of_float ?k \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 668 | by (auto simp add: \<open>0 < k\<close> \<open>1 \<le> k\<close> less_imp_le | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 669 | intro!: mult_le_one order_trans[OF _ rapprox_rat] rapprox_rat_le1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 670 | have "1 / k \<le> ?k" using rapprox_rat[where x=1 and y=k] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 671 | hence "arctan (1 / k) \<le> arctan ?k" by (rule arctan_monotone') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 672 | also have "\<dots> \<le> (?k * ub_arctan_horner prec (get_odd n) 1 ?kl)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 673 | using arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float ?k\<close> \<open>real_of_float ?k \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 674 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 675 | finally have "arctan (1 / k) \<le> ?k * ub_arctan_horner prec (get_odd n) 1 ?kl" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 676 | } note ub_arctan = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 677 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 678 |   {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 679 | fix prec n :: nat | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 680 | fix k :: int | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 681 | assume "1 < k" hence "0 \<le> k" and "0 < k" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 682 | let ?k = "lapprox_rat prec 1 k" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 683 | let ?ku = "float_round_up (Suc prec) (?k * ?k)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 684 | have "1 div k = 0" using div_pos_pos_trivial[OF _ \<open>1 < k\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 685 | have "1 / k \<le> 1" using \<open>1 < k\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 686 | have "0 \<le> real_of_float ?k" using lapprox_rat_nonneg[where x=1 and y=k, OF zero_le_one \<open>0 \<le> k\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 687 | by (auto simp add: \<open>1 div k = 0\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 688 | have "0 \<le> real_of_float (?k * ?k)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 689 | have "real_of_float ?k \<le> 1" using lapprox_rat by (rule order_trans, auto simp add: \<open>1 / k \<le> 1\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 690 | hence "real_of_float (?k * ?k) \<le> 1" using \<open>0 \<le> real_of_float ?k\<close> by (auto intro!: mult_le_one) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 691 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 692 | have "?k \<le> 1 / k" using lapprox_rat[where x=1 and y=k] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 693 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 694 | have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \<le> arctan ?k" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 695 | using arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float ?k\<close> \<open>real_of_float ?k \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 696 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 697 | also have "\<dots> \<le> arctan (1 / k)" using \<open>?k \<le> 1 / k\<close> by (rule arctan_monotone') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 698 | finally have "?k * lb_arctan_horner prec (get_even n) 1 ?ku \<le> arctan (1 / k)" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 699 | } note lb_arctan = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 700 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 701 | have "pi \<le> ub_pi n " | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 702 | unfolding ub_pi_def machin_pi Let_def times_float.rep_eq Float_num | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 703 | using lb_arctan[of 239] ub_arctan[of 5] powr_realpow[of 2 2] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 704 | by (intro mult_left_mono float_plus_up_le float_plus_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 705 | (auto intro!: mult_left_mono float_round_down_le float_round_up_le diff_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 706 | moreover have "lb_pi n \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 707 | unfolding lb_pi_def machin_pi Let_def times_float.rep_eq Float_num | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 708 | using lb_arctan[of 5] ub_arctan[of 239] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 709 | by (intro mult_left_mono float_plus_up_le float_plus_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 710 | (auto intro!: mult_left_mono float_round_down_le float_round_up_le diff_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 711 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 712 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 713 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 714 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 715 | subsection "Compute arcus tangens in the entire domain" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 716 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 717 | function lb_arctan :: "nat \<Rightarrow> float \<Rightarrow> float" and ub_arctan :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 718 | "lb_arctan prec x = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 719 | (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 720 | ub_horner = \<lambda> x. float_round_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 721 | (x * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 722 | ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))); | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 723 | lb_horner = \<lambda> x. float_round_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 724 | (x * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 725 | lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 726 | in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 727 | if x < 0 then - ub_arctan prec (-x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 728 | else if x \<le> Float 1 (- 1) then lb_horner x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 729 | else if x \<le> Float 1 1 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 730 | Float 1 1 * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 731 | lb_horner | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 732 | (float_divl prec x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 733 | (float_plus_up prec 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 734 | (ub_sqrt prec (float_plus_up prec 1 (float_round_up prec (x * x)))))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 735 | else let inv = float_divr prec 1 x in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 736 | if inv > 1 then 0 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 737 | else float_plus_down prec (lb_pi prec * Float 1 (- 1)) ( - ub_horner inv))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 738 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 739 | | "ub_arctan prec x = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 740 | (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 741 | lb_horner = \<lambda> x. float_round_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 742 | (x * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 743 | lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 744 | ub_horner = \<lambda> x. float_round_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 745 | (x * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 746 | ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 747 | in if x < 0 then - lb_arctan prec (-x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 748 | else if x \<le> Float 1 (- 1) then ub_horner x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 749 | else if x \<le> Float 1 1 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 750 | let y = float_divr prec x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 751 | (float_plus_down | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 752 | (Suc prec) 1 (lb_sqrt prec (float_plus_down prec 1 (float_round_down prec (x * x))))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 753 | in if y > 1 then ub_pi prec * Float 1 (- 1) else Float 1 1 * ub_horner y | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 754 | else float_plus_up prec (ub_pi prec * Float 1 (- 1)) ( - lb_horner (float_divl prec 1 x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 755 | by pat_completeness auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 756 | termination | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 757 | by (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 0 then 1 else 0))", auto) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 758 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 759 | declare ub_arctan_horner.simps[simp del] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 760 | declare lb_arctan_horner.simps[simp del] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 761 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 762 | lemma lb_arctan_bound': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 763 | assumes "0 \<le> real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 764 | shows "lb_arctan prec x \<le> arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 765 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 766 | have "\<not> x < 0" and "0 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 767 | using \<open>0 \<le> real_of_float x\<close> by (auto intro!: truncate_up_le ) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 768 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 769 | let "?ub_horner x" = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 770 | "x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 771 | and "?lb_horner x" = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 772 | "x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 773 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 774 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 775 | proof (cases "x \<le> Float 1 (- 1)") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 776 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 777 | hence "real_of_float x \<le> 1" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 778 | from arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 779 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 780 | unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_P[OF True] using \<open>0 \<le> x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 781 | by (auto intro!: float_round_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 782 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 783 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 784 | hence "0 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 785 | let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 786 | let ?sxx = "float_plus_up prec 1 (float_round_up prec (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 787 | let ?fR = "float_plus_up prec 1 (ub_sqrt prec ?sxx)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 788 | let ?DIV = "float_divl prec x ?fR" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 789 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 790 | have divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 791 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 792 | have "sqrt (1 + x*x) \<le> sqrt ?sxx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 793 | by (auto simp: float_plus_up.rep_eq plus_up_def float_round_up.rep_eq intro!: truncate_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 794 | also have "\<dots> \<le> ub_sqrt prec ?sxx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 795 | using bnds_sqrt'[of ?sxx prec] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 796 | finally | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 797 | have "sqrt (1 + x*x) \<le> ub_sqrt prec ?sxx" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 798 | hence "?R \<le> ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 799 | hence "0 < ?fR" and "0 < real_of_float ?fR" using \<open>0 < ?R\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 800 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 801 | have monotone: "?DIV \<le> x / ?R" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 802 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 803 | have "?DIV \<le> real_of_float x / ?fR" by (rule float_divl) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 804 | also have "\<dots> \<le> x / ?R" by (rule divide_left_mono[OF \<open>?R \<le> ?fR\<close> \<open>0 \<le> real_of_float x\<close> mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \<open>?R \<le> real_of_float ?fR\<close>] divisor_gt0]]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 805 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 806 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 807 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 808 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 809 | proof (cases "x \<le> Float 1 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 810 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 811 | have "x \<le> sqrt (1 + x * x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 812 | using real_sqrt_sum_squares_ge2[where x=1, unfolded numeral_2_eq_2] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 813 | also note \<open>\<dots> \<le> (ub_sqrt prec ?sxx)\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 814 | finally have "real_of_float x \<le> ?fR" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 815 | by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 816 | moreover have "?DIV \<le> real_of_float x / ?fR" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 817 | by (rule float_divl) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 818 | ultimately have "real_of_float ?DIV \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 819 | unfolding divide_le_eq_1_pos[OF \<open>0 < real_of_float ?fR\<close>, symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 820 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 821 | have "0 \<le> real_of_float ?DIV" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 822 | using float_divl_lower_bound[OF \<open>0 \<le> x\<close>] \<open>0 < ?fR\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 823 | unfolding less_eq_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 824 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 825 | from arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float (?DIV)\<close> \<open>real_of_float (?DIV) \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 826 | have "Float 1 1 * ?lb_horner ?DIV \<le> 2 * arctan ?DIV" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 827 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 828 | also have "\<dots> \<le> 2 * arctan (x / ?R)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 829 | using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono arctan_monotone') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 830 | also have "2 * arctan (x / ?R) = arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 831 | using arctan_half[symmetric] unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 832 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 833 | unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 834 | if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_P[OF True] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 835 | by (auto simp: float_round_down.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 836 | intro!: order_trans[OF mult_left_mono[OF truncate_down]]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 837 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 838 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 839 | hence "2 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 840 | hence "1 \<le> real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 841 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 842 | let "?invx" = "float_divr prec 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 843 | have "0 \<le> arctan x" using arctan_monotone'[OF \<open>0 \<le> real_of_float x\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 844 | using arctan_tan[of 0, unfolded tan_zero] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 845 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 846 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 847 | proof (cases "1 < ?invx") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 848 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 849 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 850 | unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 851 | if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_not_P[OF False] if_P[OF True] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 852 | using \<open>0 \<le> arctan x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 853 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 854 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 855 | hence "real_of_float ?invx \<le> 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 856 | have "0 \<le> real_of_float ?invx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 857 | by (rule order_trans[OF _ float_divr]) (auto simp add: \<open>0 \<le> real_of_float x\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 858 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 859 | have "1 / x \<noteq> 0" and "0 < 1 / x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 860 | using \<open>0 < real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 861 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 862 | have "arctan (1 / x) \<le> arctan ?invx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 863 | unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone', rule float_divr) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 864 | also have "\<dots> \<le> ?ub_horner ?invx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 865 | using arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float ?invx\<close> \<open>real_of_float ?invx \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 866 | by (auto intro!: float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 867 | also note float_round_up | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 868 | finally have "pi / 2 - float_round_up prec (?ub_horner ?invx) \<le> arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 869 | using \<open>0 \<le> arctan x\<close> arctan_inverse[OF \<open>1 / x \<noteq> 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 870 | unfolding sgn_pos[OF \<open>0 < 1 / real_of_float x\<close>] le_diff_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 871 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 872 | have "lb_pi prec * Float 1 (- 1) \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 873 | unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 874 | ultimately | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 875 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 876 | unfolding lb_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 877 | if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_not_P[OF \<open>\<not> x \<le> Float 1 1\<close>] if_not_P[OF False] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 878 | by (auto intro!: float_plus_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 879 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 880 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 881 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 882 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 883 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 884 | lemma ub_arctan_bound': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 885 | assumes "0 \<le> real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 886 | shows "arctan x \<le> ub_arctan prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 887 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 888 | have "\<not> x < 0" and "0 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 889 | using \<open>0 \<le> real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 890 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 891 | let "?ub_horner x" = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 892 | "float_round_up prec (x * ub_arctan_horner prec (get_odd (prec div 4 + 1)) 1 (float_round_down (Suc prec) (x * x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 893 | let "?lb_horner x" = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 894 | "float_round_down prec (x * lb_arctan_horner prec (get_even (prec div 4 + 1)) 1 (float_round_up (Suc prec) (x * x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 895 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 896 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 897 | proof (cases "x \<le> Float 1 (- 1)") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 898 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 899 | hence "real_of_float x \<le> 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 900 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 901 | unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] if_P[OF True] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 902 | using arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 903 | by (auto intro!: float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 904 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 905 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 906 | hence "0 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 907 | let ?R = "1 + sqrt (1 + real_of_float x * real_of_float x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 908 | let ?sxx = "float_plus_down prec 1 (float_round_down prec (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 909 | let ?fR = "float_plus_down (Suc prec) 1 (lb_sqrt prec ?sxx)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 910 | let ?DIV = "float_divr prec x ?fR" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 911 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 912 | have sqr_ge0: "0 \<le> 1 + real_of_float x * real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 913 | using sum_power2_ge_zero[of 1 "real_of_float x", unfolded numeral_2_eq_2] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 914 | hence "0 \<le> real_of_float (1 + x*x)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 915 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 916 | hence divisor_gt0: "0 < ?R" by (auto intro: add_pos_nonneg) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 917 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 918 | have "lb_sqrt prec ?sxx \<le> sqrt ?sxx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 919 | using bnds_sqrt'[of ?sxx] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 920 | also have "\<dots> \<le> sqrt (1 + x*x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 921 | by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq truncate_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 922 | finally have "lb_sqrt prec ?sxx \<le> sqrt (1 + x*x)" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 923 | hence "?fR \<le> ?R" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 924 | by (auto simp: float_plus_down.rep_eq plus_down_def truncate_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 925 | have "0 < real_of_float ?fR" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 926 | by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 927 | intro!: truncate_down_ge1 lb_sqrt_lower_bound order_less_le_trans[OF zero_less_one] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 928 | truncate_down_nonneg add_nonneg_nonneg) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 929 | have monotone: "x / ?R \<le> (float_divr prec x ?fR)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 930 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 931 | from divide_left_mono[OF \<open>?fR \<le> ?R\<close> \<open>0 \<le> real_of_float x\<close> mult_pos_pos[OF divisor_gt0 \<open>0 < real_of_float ?fR\<close>]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 932 | have "x / ?R \<le> x / ?fR" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 933 | also have "\<dots> \<le> ?DIV" by (rule float_divr) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 934 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 935 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 936 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 937 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 938 | proof (cases "x \<le> Float 1 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 939 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 940 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 941 | proof (cases "?DIV > 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 942 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 943 | have "pi / 2 \<le> ub_pi prec * Float 1 (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 944 | unfolding Float_num times_divide_eq_right mult_1_left using pi_boundaries by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 945 | from order_less_le_trans[OF arctan_ubound this, THEN less_imp_le] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 946 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 947 | unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 948 | if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_P[OF \<open>x \<le> Float 1 1\<close>] if_P[OF True] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 949 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 950 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 951 | hence "real_of_float ?DIV \<le> 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 952 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 953 | have "0 \<le> x / ?R" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 954 | using \<open>0 \<le> real_of_float x\<close> \<open>0 < ?R\<close> unfolding zero_le_divide_iff by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 955 | hence "0 \<le> real_of_float ?DIV" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 956 | using monotone by (rule order_trans) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 957 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 958 | have "arctan x = 2 * arctan (x / ?R)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 959 | using arctan_half unfolding numeral_2_eq_2 power_Suc2 power_0 mult_1_left . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 960 | also have "\<dots> \<le> 2 * arctan (?DIV)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 961 | using arctan_monotone'[OF monotone] by (auto intro!: mult_left_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 962 | also have "\<dots> \<le> (Float 1 1 * ?ub_horner ?DIV)" unfolding Float_num | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 963 | using arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float ?DIV\<close> \<open>real_of_float ?DIV \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 964 | by (auto intro!: float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 965 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 966 | unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 967 | if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_P[OF \<open>x \<le> Float 1 1\<close>] if_not_P[OF False] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 968 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 969 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 970 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 971 | hence "2 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 972 | hence "1 \<le> real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 973 | hence "0 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 974 | hence "0 < x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 975 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 976 | let "?invx" = "float_divl prec 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 977 | have "0 \<le> arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 978 | using arctan_monotone'[OF \<open>0 \<le> real_of_float x\<close>] and arctan_tan[of 0, unfolded tan_zero] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 979 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 980 | have "real_of_float ?invx \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 981 | unfolding less_float_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 982 | by (rule order_trans[OF float_divl]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 983 | (auto simp add: \<open>1 \<le> real_of_float x\<close> divide_le_eq_1_pos[OF \<open>0 < real_of_float x\<close>]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 984 | have "0 \<le> real_of_float ?invx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 985 | using \<open>0 < x\<close> by (intro float_divl_lower_bound) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 986 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 987 | have "1 / x \<noteq> 0" and "0 < 1 / x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 988 | using \<open>0 < real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 989 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 990 | have "(?lb_horner ?invx) \<le> arctan (?invx)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 991 | using arctan_0_1_bounds_round[OF \<open>0 \<le> real_of_float ?invx\<close> \<open>real_of_float ?invx \<le> 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 992 | by (auto intro!: float_round_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 993 | also have "\<dots> \<le> arctan (1 / x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 994 | unfolding one_float.rep_eq[symmetric] by (rule arctan_monotone') (rule float_divl) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 995 | finally have "arctan x \<le> pi / 2 - (?lb_horner ?invx)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 996 | using \<open>0 \<le> arctan x\<close> arctan_inverse[OF \<open>1 / x \<noteq> 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 997 | unfolding sgn_pos[OF \<open>0 < 1 / x\<close>] le_diff_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 998 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 999 | have "pi / 2 \<le> ub_pi prec * Float 1 (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1000 | unfolding Float_num times_divide_eq_right mult_1_right | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1001 | using pi_boundaries by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1002 | ultimately | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1003 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1004 | unfolding ub_arctan.simps Let_def if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1005 | if_not_P[OF \<open>\<not> x \<le> Float 1 (- 1)\<close>] if_not_P[OF False] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1006 | by (auto intro!: float_round_up_le float_plus_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1007 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1008 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1009 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1010 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1011 | lemma arctan_boundaries: "arctan x \<in> {(lb_arctan prec x) .. (ub_arctan prec x)}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1012 | proof (cases "0 \<le> x") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1013 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1014 | hence "0 \<le> real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1015 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1016 | using ub_arctan_bound'[OF \<open>0 \<le> real_of_float x\<close>] lb_arctan_bound'[OF \<open>0 \<le> real_of_float x\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1017 | unfolding atLeastAtMost_iff by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1018 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1019 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1020 | let ?mx = "-x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1021 | from False have "x < 0" and "0 \<le> real_of_float ?mx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1022 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1023 | hence bounds: "lb_arctan prec ?mx \<le> arctan ?mx \<and> arctan ?mx \<le> ub_arctan prec ?mx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1024 | using ub_arctan_bound'[OF \<open>0 \<le> real_of_float ?mx\<close>] lb_arctan_bound'[OF \<open>0 \<le> real_of_float ?mx\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1025 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1026 | unfolding minus_float.rep_eq arctan_minus lb_arctan.simps[where x=x] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1027 | ub_arctan.simps[where x=x] Let_def if_P[OF \<open>x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1028 | unfolding atLeastAtMost_iff using bounds[unfolded minus_float.rep_eq arctan_minus] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1029 | by (simp add: arctan_minus) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1030 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1031 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1032 | lemma bnds_arctan: "\<forall> (x::real) lx ux. (l, u) = (lb_arctan prec lx, ub_arctan prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> arctan x \<and> arctan x \<le> u"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1033 | proof (rule allI, rule allI, rule allI, rule impI) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1034 | fix x :: real | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1035 | fix lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1036 |   assume "(l, u) = (lb_arctan prec lx, ub_arctan prec ux) \<and> x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1037 | hence l: "lb_arctan prec lx = l " | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1038 | and u: "ub_arctan prec ux = u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1039 |     and x: "x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1040 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1041 | show "l \<le> arctan x \<and> arctan x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1042 | proof | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1043 | show "l \<le> arctan x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1044 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1045 | from arctan_boundaries[of lx prec, unfolded l] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1046 | have "l \<le> arctan lx" by (auto simp del: lb_arctan.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1047 | also have "\<dots> \<le> arctan x" using x by (auto intro: arctan_monotone') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1048 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1049 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1050 | show "arctan x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1051 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1052 | have "arctan x \<le> arctan ux" using x by (auto intro: arctan_monotone') | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1053 | also have "\<dots> \<le> u" using arctan_boundaries[of ux prec, unfolded u] by (auto simp del: ub_arctan.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1054 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1055 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1056 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1057 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1058 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1059 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1060 | section "Sinus and Cosinus" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1061 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1062 | subsection "Compute the cosinus and sinus series" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1063 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1064 | fun ub_sin_cos_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1065 | and lb_sin_cos_aux :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1066 | "ub_sin_cos_aux prec 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1067 | | "ub_sin_cos_aux prec (Suc n) i k x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1068 | (rapprox_rat prec 1 k) (- | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1069 | float_round_down prec (x * (lb_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1070 | | "lb_sin_cos_aux prec 0 i k x = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1071 | | "lb_sin_cos_aux prec (Suc n) i k x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1072 | (lapprox_rat prec 1 k) (- | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1073 | float_round_up prec (x * (ub_sin_cos_aux prec n (i + 2) (k * i * (i + 1)) x)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1074 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1075 | lemma cos_aux: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1076 | shows "(lb_sin_cos_aux prec n 1 1 (x * x)) \<le> (\<Sum> i=0..<n. (- 1) ^ i * (1/(fact (2 * i))) * x ^(2 * i))" (is "?lb") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1077 | and "(\<Sum> i=0..<n. (- 1) ^ i * (1/(fact (2 * i))) * x^(2 * i)) \<le> (ub_sin_cos_aux prec n 1 1 (x * x))" (is "?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1078 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1079 | have "0 \<le> real_of_float (x * x)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1080 | let "?f n" = "fact (2 * n) :: nat" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1081 | have f_eq: "?f (Suc n) = ?f n * ((\<lambda>i. i + 2) ^^ n) 1 * (((\<lambda>i. i + 2) ^^ n) 1 + 1)" for n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1082 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1083 | have "\<And>m. ((\<lambda>i. i + 2) ^^ n) m = m + 2 * n" by (induct n) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1084 | then show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1085 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1086 | from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1087 | OF \<open>0 \<le> real_of_float (x * x)\<close> f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1088 | show ?lb and ?ub | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1089 | by (auto simp add: power_mult power2_eq_square[of "real_of_float x"]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1090 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1091 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1092 | lemma lb_sin_cos_aux_zero_le_one: "lb_sin_cos_aux prec n i j 0 \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1093 | by (cases j n rule: nat.exhaust[case_product nat.exhaust]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1094 | (auto intro!: float_plus_down_le order_trans[OF lapprox_rat]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1095 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1096 | lemma one_le_ub_sin_cos_aux: "odd n \<Longrightarrow> 1 \<le> ub_sin_cos_aux prec n i (Suc 0) 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1097 | by (cases n) (auto intro!: float_plus_up_le order_trans[OF _ rapprox_rat]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1098 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1099 | lemma cos_boundaries: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1100 | assumes "0 \<le> real_of_float x" and "x \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1101 |   shows "cos x \<in> {(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) .. (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1102 | proof (cases "real_of_float x = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1103 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1104 | hence "real_of_float x \<noteq> 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1105 | hence "0 < x" and "0 < real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1106 | using \<open>0 \<le> real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1107 | have "0 < x * x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1108 | using \<open>0 < x\<close> by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1109 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1110 | have morph_to_if_power: "(\<Sum> i=0..<n. (-1::real) ^ i * (1/(fact (2 * i))) * x ^ (2 * i)) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1111 | (\<Sum> i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * x ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1112 | (is "?sum = ?ifsum") for x n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1113 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1114 | have "?sum = ?sum + (\<Sum> j = 0 ..< n. 0)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1115 | also have "\<dots> = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1116 | (\<Sum> j = 0 ..< n. (- 1) ^ ((2 * j) div 2) / ((fact (2 * j))) * x ^(2 * j)) + (\<Sum> j = 0 ..< n. 0)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1117 | also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then (- 1) ^ (i div 2) / ((fact i)) * x ^ i else 0)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1118 | unfolding sum_split_even_odd atLeast0LessThan .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1119 | also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then (- 1) ^ (i div 2) / ((fact i)) else 0) * x ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1120 | by (rule sum.cong) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1121 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1122 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1123 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1124 |   { fix n :: nat assume "0 < n"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1125 | hence "0 < 2 * n" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1126 | obtain t where "0 < t" and "t < real_of_float x" and | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1127 | cos_eq: "cos x = (\<Sum> i = 0 ..< 2 * n. (if even(i) then ((- 1) ^ (i div 2))/((fact i)) else 0) * (real_of_float x) ^ i) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1128 | + (cos (t + 1/2 * (2 * n) * pi) / (fact (2*n))) * (real_of_float x)^(2*n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1129 | (is "_ = ?SUM + ?rest / ?fact * ?pow") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1130 | using Maclaurin_cos_expansion2[OF \<open>0 < real_of_float x\<close> \<open>0 < 2 * n\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1131 | unfolding cos_coeff_def atLeast0LessThan by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1132 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1133 | have "cos t * (- 1) ^ n = cos t * cos (n * pi) + sin t * sin (n * pi)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1134 | also have "\<dots> = cos (t + n * pi)" by (simp add: cos_add) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1135 | also have "\<dots> = ?rest" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1136 | finally have "cos t * (- 1) ^ n = ?rest" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1137 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1138 | have "t \<le> pi / 2" using \<open>t < real_of_float x\<close> and \<open>x \<le> pi / 2\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1139 | hence "0 \<le> cos t" using \<open>0 < t\<close> and cos_ge_zero by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1140 | ultimately have even: "even n \<Longrightarrow> 0 \<le> ?rest" and odd: "odd n \<Longrightarrow> 0 \<le> - ?rest " by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1141 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1142 | have "0 < ?fact" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1143 | have "0 < ?pow" using \<open>0 < real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1144 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1145 |     {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1146 | assume "even n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1147 | have "(lb_sin_cos_aux prec n 1 1 (x * x)) \<le> ?SUM" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1148 | unfolding morph_to_if_power[symmetric] using cos_aux by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1149 | also have "\<dots> \<le> cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1150 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1151 | from even[OF \<open>even n\<close>] \<open>0 < ?fact\<close> \<open>0 < ?pow\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1152 | have "0 \<le> (?rest / ?fact) * ?pow" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1153 | thus ?thesis unfolding cos_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1154 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1155 | finally have "(lb_sin_cos_aux prec n 1 1 (x * x)) \<le> cos x" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1156 | } note lb = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1157 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1158 |     {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1159 | assume "odd n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1160 | have "cos x \<le> ?SUM" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1161 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1162 | from \<open>0 < ?fact\<close> and \<open>0 < ?pow\<close> and odd[OF \<open>odd n\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1163 | have "0 \<le> (- ?rest) / ?fact * ?pow" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1164 | by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1165 | thus ?thesis unfolding cos_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1166 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1167 | also have "\<dots> \<le> (ub_sin_cos_aux prec n 1 1 (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1168 | unfolding morph_to_if_power[symmetric] using cos_aux by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1169 | finally have "cos x \<le> (ub_sin_cos_aux prec n 1 1 (x * x))" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1170 | } note ub = this and lb | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1171 | } note ub = this(1) and lb = this(2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1172 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1173 | have "cos x \<le> (ub_sin_cos_aux prec (get_odd n) 1 1 (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1174 | using ub[OF odd_pos[OF get_odd] get_odd] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1175 | moreover have "(lb_sin_cos_aux prec (get_even n) 1 1 (x * x)) \<le> cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1176 | proof (cases "0 < get_even n") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1177 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1178 | show ?thesis using lb[OF True get_even] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1179 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1180 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1181 | hence "get_even n = 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1182 | have "- (pi / 2) \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1183 | by (rule order_trans[OF _ \<open>0 < real_of_float x\<close>[THEN less_imp_le]]) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1184 | with \<open>x \<le> pi / 2\<close> show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1185 | unfolding \<open>get_even n = 0\<close> lb_sin_cos_aux.simps minus_float.rep_eq zero_float.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1186 | using cos_ge_zero by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1187 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1188 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1189 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1190 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1191 | hence "x = 0" | 
| 67573 | 1192 | by (simp add: real_of_float_eq) | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1193 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1194 | using lb_sin_cos_aux_zero_le_one one_le_ub_sin_cos_aux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1195 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1196 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1197 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1198 | lemma sin_aux: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1199 | assumes "0 \<le> real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1200 | shows "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1201 | (\<Sum> i=0..<n. (- 1) ^ i * (1/(fact (2 * i + 1))) * x^(2 * i + 1))" (is "?lb") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1202 | and "(\<Sum> i=0..<n. (- 1) ^ i * (1/(fact (2 * i + 1))) * x^(2 * i + 1)) \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1203 | (x * ub_sin_cos_aux prec n 2 1 (x * x))" (is "?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1204 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1205 | have "0 \<le> real_of_float (x * x)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1206 | let "?f n" = "fact (2 * n + 1) :: nat" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1207 | have f_eq: "?f (Suc n) = ?f n * ((\<lambda>i. i + 2) ^^ n) 2 * (((\<lambda>i. i + 2) ^^ n) 2 + 1)" for n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1208 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1209 | have F: "\<And>m. ((\<lambda>i. i + 2) ^^ n) m = m + 2 * n" by (induct n) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1210 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1211 | unfolding F by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1212 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1213 | from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0, | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1214 | OF \<open>0 \<le> real_of_float (x * x)\<close> f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1215 | show "?lb" and "?ub" using \<open>0 \<le> real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1216 | apply (simp_all only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1217 | apply (simp_all only: mult.commute[where 'a=real] of_nat_fact) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1218 | apply (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real_of_float x"]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1219 | done | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1220 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1221 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1222 | lemma sin_boundaries: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1223 | assumes "0 \<le> real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1224 | and "x \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1225 |   shows "sin x \<in> {(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) .. (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1226 | proof (cases "real_of_float x = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1227 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1228 | hence "real_of_float x \<noteq> 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1229 | hence "0 < x" and "0 < real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1230 | using \<open>0 \<le> real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1231 | have "0 < x * x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1232 | using \<open>0 < x\<close> by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1233 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1234 | have sum_morph: "(\<Sum>j = 0 ..< n. (- 1) ^ (((2 * j + 1) - Suc 0) div 2) / ((fact (2 * j + 1))) * x ^(2 * j + 1)) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1235 | (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * x ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1236 | (is "?SUM = _") for x :: real and n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1237 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1238 | have pow: "!!i. x ^ (2 * i + 1) = x * x ^ (2 * i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1239 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1240 | have "?SUM = (\<Sum> j = 0 ..< n. 0) + ?SUM" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1241 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1242 | also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i)) * x ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1243 | unfolding sum_split_even_odd atLeast0LessThan .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1244 | also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i))) * x ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1245 | by (rule sum.cong) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1246 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1247 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1248 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1249 |   { fix n :: nat assume "0 < n"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1250 | hence "0 < 2 * n + 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1251 | obtain t where "0 < t" and "t < real_of_float x" and | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1252 | sin_eq: "sin x = (\<Sum> i = 0 ..< 2 * n + 1. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1253 | + (sin (t + 1/2 * (2 * n + 1) * pi) / (fact (2*n + 1))) * (real_of_float x)^(2*n + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1254 | (is "_ = ?SUM + ?rest / ?fact * ?pow") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1255 | using Maclaurin_sin_expansion3[OF \<open>0 < 2 * n + 1\<close> \<open>0 < real_of_float x\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1256 | unfolding sin_coeff_def atLeast0LessThan by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1257 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1258 | have "?rest = cos t * (- 1) ^ n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1259 | unfolding sin_add cos_add of_nat_add distrib_right distrib_left by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1260 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1261 | have "t \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1262 | using \<open>t < real_of_float x\<close> and \<open>x \<le> pi / 2\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1263 | hence "0 \<le> cos t" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1264 | using \<open>0 < t\<close> and cos_ge_zero by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1265 | ultimately have even: "even n \<Longrightarrow> 0 \<le> ?rest" and odd: "odd n \<Longrightarrow> 0 \<le> - ?rest" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1266 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1267 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1268 | have "0 < ?fact" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1269 | by (simp del: fact_Suc) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1270 | have "0 < ?pow" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1271 | using \<open>0 < real_of_float x\<close> by (rule zero_less_power) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1272 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1273 |     {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1274 | assume "even n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1275 | have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1276 | (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1277 | using sin_aux[OF \<open>0 \<le> real_of_float x\<close>] unfolding sum_morph[symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1278 | also have "\<dots> \<le> ?SUM" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1279 | also have "\<dots> \<le> sin x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1280 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1281 | from even[OF \<open>even n\<close>] \<open>0 < ?fact\<close> \<open>0 < ?pow\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1282 | have "0 \<le> (?rest / ?fact) * ?pow" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1283 | thus ?thesis unfolding sin_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1284 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1285 | finally have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le> sin x" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1286 | } note lb = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1287 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1288 |     {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1289 | assume "odd n" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1290 | have "sin x \<le> ?SUM" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1291 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1292 | from \<open>0 < ?fact\<close> and \<open>0 < ?pow\<close> and odd[OF \<open>odd n\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1293 | have "0 \<le> (- ?rest) / ?fact * ?pow" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1294 | by (metis mult_nonneg_nonneg divide_nonneg_pos less_imp_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1295 | thus ?thesis unfolding sin_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1296 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1297 | also have "\<dots> \<le> (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1298 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1299 | also have "\<dots> \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1300 | using sin_aux[OF \<open>0 \<le> real_of_float x\<close>] unfolding sum_morph[symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1301 | finally have "sin x \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1302 | } note ub = this and lb | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1303 | } note ub = this(1) and lb = this(2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1304 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1305 | have "sin x \<le> (x * ub_sin_cos_aux prec (get_odd n) 2 1 (x * x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1306 | using ub[OF odd_pos[OF get_odd] get_odd] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1307 | moreover have "(x * lb_sin_cos_aux prec (get_even n) 2 1 (x * x)) \<le> sin x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1308 | proof (cases "0 < get_even n") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1309 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1310 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1311 | using lb[OF True get_even] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1312 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1313 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1314 | hence "get_even n = 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1315 | with \<open>x \<le> pi / 2\<close> \<open>0 \<le> real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1316 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1317 | unfolding \<open>get_even n = 0\<close> ub_sin_cos_aux.simps minus_float.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1318 | using sin_ge_zero by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1319 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1320 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1321 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1322 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1323 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1324 | proof (cases "n = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1325 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1326 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1327 | unfolding \<open>n = 0\<close> get_even_def get_odd_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1328 | using \<open>real_of_float x = 0\<close> lapprox_rat[where x="-1" and y=1] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1329 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1330 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1331 | with not0_implies_Suc obtain m where "n = Suc m" by blast | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1332 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1333 | unfolding \<open>n = Suc m\<close> get_even_def get_odd_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1334 | using \<open>real_of_float x = 0\<close> rapprox_rat[where x=1 and y=1] lapprox_rat[where x=1 and y=1] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1335 | by (cases "even (Suc m)") auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1336 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1337 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1338 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1339 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1340 | subsection "Compute the cosinus in the entire domain" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1341 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1342 | definition lb_cos :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1343 | "lb_cos prec x = (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1344 | horner = \<lambda> x. lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1345 | half = \<lambda> x. if x < 0 then - 1 else float_plus_down prec (Float 1 1 * x * x) (- 1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1346 | in if x < Float 1 (- 1) then horner x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1347 | else if x < 1 then half (horner (x * Float 1 (- 1))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1348 | else half (half (horner (x * Float 1 (- 2)))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1349 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1350 | definition ub_cos :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1351 | "ub_cos prec x = (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1352 | horner = \<lambda> x. ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1353 | half = \<lambda> x. float_plus_up prec (Float 1 1 * x * x) (- 1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1354 | in if x < Float 1 (- 1) then horner x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1355 | else if x < 1 then half (horner (x * Float 1 (- 1))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1356 | else half (half (horner (x * Float 1 (- 2)))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1357 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1358 | lemma lb_cos: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1359 | assumes "0 \<le> real_of_float x" and "x \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1360 |   shows "cos x \<in> {(lb_cos prec x) .. (ub_cos prec x)}" (is "?cos x \<in> {(?lb x) .. (?ub x) }")
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1361 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1362 | have x_half[symmetric]: "cos x = 2 * cos (x / 2) * cos (x / 2) - 1" for x :: real | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1363 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1364 | have "cos x = cos (x / 2 + x / 2)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1365 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1366 | also have "\<dots> = cos (x / 2) * cos (x / 2) + sin (x / 2) * sin (x / 2) - sin (x / 2) * sin (x / 2) + cos (x / 2) * cos (x / 2) - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1367 | unfolding cos_add by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1368 | also have "\<dots> = 2 * cos (x / 2) * cos (x / 2) - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1369 | by algebra | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1370 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1371 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1372 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1373 | have "\<not> x < 0" using \<open>0 \<le> real_of_float x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1374 | let "?ub_horner x" = "ub_sin_cos_aux prec (get_odd (prec div 4 + 1)) 1 1 (x * x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1375 | let "?lb_horner x" = "lb_sin_cos_aux prec (get_even (prec div 4 + 1)) 1 1 (x * x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1376 | let "?ub_half x" = "float_plus_up prec (Float 1 1 * x * x) (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1377 | let "?lb_half x" = "if x < 0 then - 1 else float_plus_down prec (Float 1 1 * x * x) (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1378 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1379 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1380 | proof (cases "x < Float 1 (- 1)") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1381 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1382 | hence "x \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1383 | using pi_ge_two by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1384 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1385 | unfolding lb_cos_def[where x=x] ub_cos_def[where x=x] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1386 | if_not_P[OF \<open>\<not> x < 0\<close>] if_P[OF \<open>x < Float 1 (- 1)\<close>] Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1387 | using cos_boundaries[OF \<open>0 \<le> real_of_float x\<close> \<open>x \<le> pi / 2\<close>] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1388 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1389 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1390 |     { fix y x :: float let ?x2 = "(x * Float 1 (- 1))"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1391 | assume "y \<le> cos ?x2" and "-pi \<le> x" and "x \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1392 | hence "- (pi / 2) \<le> ?x2" and "?x2 \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1393 | using pi_ge_two unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1394 | hence "0 \<le> cos ?x2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1395 | by (rule cos_ge_zero) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1396 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1397 | have "(?lb_half y) \<le> cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1398 | proof (cases "y < 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1399 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1400 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1401 | using cos_ge_minus_one unfolding if_P[OF True] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1402 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1403 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1404 | hence "0 \<le> real_of_float y" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1405 | from mult_mono[OF \<open>y \<le> cos ?x2\<close> \<open>y \<le> cos ?x2\<close> \<open>0 \<le> cos ?x2\<close> this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1406 | have "real_of_float y * real_of_float y \<le> cos ?x2 * cos ?x2" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1407 | hence "2 * real_of_float y * real_of_float y \<le> 2 * cos ?x2 * cos ?x2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1408 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1409 | hence "2 * real_of_float y * real_of_float y - 1 \<le> 2 * cos (x / 2) * cos (x / 2) - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1410 | unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1411 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1412 | unfolding if_not_P[OF False] x_half Float_num | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1413 | by (auto intro!: float_plus_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1414 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1415 | } note lb_half = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1416 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1417 |     { fix y x :: float let ?x2 = "(x * Float 1 (- 1))"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1418 | assume ub: "cos ?x2 \<le> y" and "- pi \<le> x" and "x \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1419 | hence "- (pi / 2) \<le> ?x2" and "?x2 \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1420 | using pi_ge_two unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1421 | hence "0 \<le> cos ?x2" by (rule cos_ge_zero) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1422 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1423 | have "cos x \<le> (?ub_half y)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1424 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1425 | have "0 \<le> real_of_float y" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1426 | using \<open>0 \<le> cos ?x2\<close> ub by (rule order_trans) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1427 | from mult_mono[OF ub ub this \<open>0 \<le> cos ?x2\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1428 | have "cos ?x2 * cos ?x2 \<le> real_of_float y * real_of_float y" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1429 | hence "2 * cos ?x2 * cos ?x2 \<le> 2 * real_of_float y * real_of_float y" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1430 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1431 | hence "2 * cos (x / 2) * cos (x / 2) - 1 \<le> 2 * real_of_float y * real_of_float y - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1432 | unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1433 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1434 | unfolding x_half Float_num | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1435 | by (auto intro!: float_plus_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1436 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1437 | } note ub_half = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1438 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1439 | let ?x2 = "x * Float 1 (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1440 | let ?x4 = "x * Float 1 (- 1) * Float 1 (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1441 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1442 | have "-pi \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1443 | using pi_ge_zero[THEN le_imp_neg_le, unfolded minus_zero] \<open>0 \<le> real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1444 | by (rule order_trans) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1445 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1446 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1447 | proof (cases "x < 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1448 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1449 | hence "real_of_float x \<le> 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1450 | have "0 \<le> real_of_float ?x2" and "?x2 \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1451 | using pi_ge_two \<open>0 \<le> real_of_float x\<close> using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1452 | from cos_boundaries[OF this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1453 | have lb: "(?lb_horner ?x2) \<le> ?cos ?x2" and ub: "?cos ?x2 \<le> (?ub_horner ?x2)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1454 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1455 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1456 | have "(?lb x) \<le> ?cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1457 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1458 | from lb_half[OF lb \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1459 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1460 | unfolding lb_cos_def[where x=x] Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1461 | using \<open>\<not> x < 0\<close> \<open>\<not> x < Float 1 (- 1)\<close> \<open>x < 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1462 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1463 | moreover have "?cos x \<le> (?ub x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1464 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1465 | from ub_half[OF ub \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1466 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1467 | unfolding ub_cos_def[where x=x] Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1468 | using \<open>\<not> x < 0\<close> \<open>\<not> x < Float 1 (- 1)\<close> \<open>x < 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1469 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1470 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1471 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1472 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1473 | have "0 \<le> real_of_float ?x4" and "?x4 \<le> pi / 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1474 | using pi_ge_two \<open>0 \<le> real_of_float x\<close> \<open>x \<le> pi\<close> unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1475 | from cos_boundaries[OF this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1476 | have lb: "(?lb_horner ?x4) \<le> ?cos ?x4" and ub: "?cos ?x4 \<le> (?ub_horner ?x4)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1477 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1478 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1479 | have eq_4: "?x2 * Float 1 (- 1) = x * Float 1 (- 2)" | 
| 67573 | 1480 | by (auto simp: real_of_float_eq) | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1481 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1482 | have "(?lb x) \<le> ?cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1483 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1484 | have "-pi \<le> ?x2" and "?x2 \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1485 | using pi_ge_two \<open>0 \<le> real_of_float x\<close> \<open>x \<le> pi\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1486 | from lb_half[OF lb_half[OF lb this] \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>, unfolded eq_4] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1487 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1488 | unfolding lb_cos_def[where x=x] if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1489 | if_not_P[OF \<open>\<not> x < Float 1 (- 1)\<close>] if_not_P[OF \<open>\<not> x < 1\<close>] Let_def . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1490 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1491 | moreover have "?cos x \<le> (?ub x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1492 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1493 | have "-pi \<le> ?x2" and "?x2 \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1494 | using pi_ge_two \<open>0 \<le> real_of_float x\<close> \<open> x \<le> pi\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1495 | from ub_half[OF ub_half[OF ub this] \<open>-pi \<le> x\<close> \<open>x \<le> pi\<close>, unfolded eq_4] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1496 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1497 | unfolding ub_cos_def[where x=x] if_not_P[OF \<open>\<not> x < 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1498 | if_not_P[OF \<open>\<not> x < Float 1 (- 1)\<close>] if_not_P[OF \<open>\<not> x < 1\<close>] Let_def . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1499 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1500 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1501 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1502 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1503 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1504 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1505 | lemma lb_cos_minus: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1506 | assumes "-pi \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1507 | and "real_of_float x \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1508 |   shows "cos (real_of_float(-x)) \<in> {(lb_cos prec (-x)) .. (ub_cos prec (-x))}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1509 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1510 | have "0 \<le> real_of_float (-x)" and "(-x) \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1511 | using \<open>-pi \<le> x\<close> \<open>real_of_float x \<le> 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1512 | from lb_cos[OF this] show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1513 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1514 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1515 | definition bnds_cos :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float * float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1516 | "bnds_cos prec lx ux = (let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1517 | lpi = float_round_down prec (lb_pi prec) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1518 | upi = float_round_up prec (ub_pi prec) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1519 | k = floor_fl (float_divr prec (lx + lpi) (2 * lpi)) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1520 | lx = float_plus_down prec lx (- k * 2 * (if k < 0 then lpi else upi)) ; | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1521 | ux = float_plus_up prec ux (- k * 2 * (if k < 0 then upi else lpi)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1522 | in if - lpi \<le> lx \<and> ux \<le> 0 then (lb_cos prec (-lx), ub_cos prec (-ux)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1523 | else if 0 \<le> lx \<and> ux \<le> lpi then (lb_cos prec ux, ub_cos prec lx) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1524 | else if - lpi \<le> lx \<and> ux \<le> lpi then (min (lb_cos prec (-lx)) (lb_cos prec ux), Float 1 0) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1525 | else if 0 \<le> lx \<and> ux \<le> 2 * lpi then (Float (- 1) 0, max (ub_cos prec lx) (ub_cos prec (- (ux - 2 * lpi)))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1526 | else if -2 * lpi \<le> lx \<and> ux \<le> 0 then (Float (- 1) 0, max (ub_cos prec (lx + 2 * lpi)) (ub_cos prec (-ux))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1527 | else (Float (- 1) 0, Float 1 0))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1528 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1529 | lemma floor_int: obtains k :: int where "real_of_int k = (floor_fl f)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1530 | by (simp add: floor_fl_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1531 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1532 | lemma cos_periodic_nat[simp]: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1533 | fixes n :: nat | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1534 | shows "cos (x + n * (2 * pi)) = cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1535 | proof (induct n arbitrary: x) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1536 | case 0 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1537 | then show ?case by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1538 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1539 | case (Suc n) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1540 | have split_pi_off: "x + (Suc n) * (2 * pi) = (x + n * (2 * pi)) + 2 * pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1541 | unfolding Suc_eq_plus1 of_nat_add of_int_1 distrib_right by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1542 | show ?case | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1543 | unfolding split_pi_off using Suc by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1544 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1545 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1546 | lemma cos_periodic_int[simp]: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1547 | fixes i :: int | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1548 | shows "cos (x + i * (2 * pi)) = cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1549 | proof (cases "0 \<le> i") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1550 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1551 | hence i_nat: "real_of_int i = nat i" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1552 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1553 | unfolding i_nat by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1554 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1555 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1556 | hence i_nat: "i = - real (nat (-i))" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1557 | have "cos x = cos (x + i * (2 * pi) - i * (2 * pi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1558 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1559 | also have "\<dots> = cos (x + i * (2 * pi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1560 | unfolding i_nat mult_minus_left diff_minus_eq_add by (rule cos_periodic_nat) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1561 | finally show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1562 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1563 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1564 | lemma bnds_cos: "\<forall>(x::real) lx ux. (l, u) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1565 |   bnds_cos prec lx ux \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> cos x \<and> cos x \<le> u"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1566 | proof (rule allI | rule impI | erule conjE)+ | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1567 | fix x :: real | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1568 | fix lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1569 |   assume bnds: "(l, u) = bnds_cos prec lx ux" and x: "x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1570 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1571 | let ?lpi = "float_round_down prec (lb_pi prec)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1572 | let ?upi = "float_round_up prec (ub_pi prec)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1573 | let ?k = "floor_fl (float_divr prec (lx + ?lpi) (2 * ?lpi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1574 | let ?lx2 = "(- ?k * 2 * (if ?k < 0 then ?lpi else ?upi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1575 | let ?ux2 = "(- ?k * 2 * (if ?k < 0 then ?upi else ?lpi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1576 | let ?lx = "float_plus_down prec lx ?lx2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1577 | let ?ux = "float_plus_up prec ux ?ux2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1578 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1579 | obtain k :: int where k: "k = real_of_float ?k" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1580 | by (rule floor_int) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1581 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1582 | have upi: "pi \<le> ?upi" and lpi: "?lpi \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1583 | using float_round_up[of "ub_pi prec" prec] pi_boundaries[of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1584 | float_round_down[of prec "lb_pi prec"] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1585 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1586 | hence "lx + ?lx2 \<le> x - k * (2 * pi) \<and> x - k * (2 * pi) \<le> ux + ?ux2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1587 | using x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1588 | by (cases "k = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1589 | (auto intro!: add_mono | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1590 | simp add: k [symmetric] uminus_add_conv_diff [symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1591 | simp del: float_of_numeral uminus_add_conv_diff) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1592 | hence "?lx \<le> x - k * (2 * pi) \<and> x - k * (2 * pi) \<le> ?ux" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1593 | by (auto intro!: float_plus_down_le float_plus_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1594 | note lx = this[THEN conjunct1] and ux = this[THEN conjunct2] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1595 | hence lx_less_ux: "?lx \<le> real_of_float ?ux" by (rule order_trans) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1596 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1597 |   { assume "- ?lpi \<le> ?lx" and x_le_0: "x - k * (2 * pi) \<le> 0"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1598 | with lpi[THEN le_imp_neg_le] lx | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1599 | have pi_lx: "- pi \<le> ?lx" and lx_0: "real_of_float ?lx \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1600 | by simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1601 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1602 | have "(lb_cos prec (- ?lx)) \<le> cos (real_of_float (- ?lx))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1603 | using lb_cos_minus[OF pi_lx lx_0] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1604 | also have "\<dots> \<le> cos (x + (-k) * (2 * pi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1605 | using cos_monotone_minus_pi_0'[OF pi_lx lx x_le_0] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1606 | by (simp only: uminus_float.rep_eq of_int_minus | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1607 | cos_minus mult_minus_left) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1608 | finally have "(lb_cos prec (- ?lx)) \<le> cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1609 | unfolding cos_periodic_int . } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1610 | note negative_lx = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1611 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1612 |   { assume "0 \<le> ?lx" and pi_x: "x - k * (2 * pi) \<le> pi"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1613 | with lx | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1614 | have pi_lx: "?lx \<le> pi" and lx_0: "0 \<le> real_of_float ?lx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1615 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1616 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1617 | have "cos (x + (-k) * (2 * pi)) \<le> cos ?lx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1618 | using cos_monotone_0_pi_le[OF lx_0 lx pi_x] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1619 | by (simp only: of_int_minus | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1620 | cos_minus mult_minus_left) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1621 | also have "\<dots> \<le> (ub_cos prec ?lx)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1622 | using lb_cos[OF lx_0 pi_lx] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1623 | finally have "cos x \<le> (ub_cos prec ?lx)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1624 | unfolding cos_periodic_int . } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1625 | note positive_lx = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1626 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1627 |   { assume pi_x: "- pi \<le> x - k * (2 * pi)" and "?ux \<le> 0"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1628 | with ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1629 | have pi_ux: "- pi \<le> ?ux" and ux_0: "real_of_float ?ux \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1630 | by simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1631 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1632 | have "cos (x + (-k) * (2 * pi)) \<le> cos (real_of_float (- ?ux))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1633 | using cos_monotone_minus_pi_0'[OF pi_x ux ux_0] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1634 | by (simp only: uminus_float.rep_eq of_int_minus | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1635 | cos_minus mult_minus_left) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1636 | also have "\<dots> \<le> (ub_cos prec (- ?ux))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1637 | using lb_cos_minus[OF pi_ux ux_0, of prec] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1638 | finally have "cos x \<le> (ub_cos prec (- ?ux))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1639 | unfolding cos_periodic_int . } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1640 | note negative_ux = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1641 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1642 |   { assume "?ux \<le> ?lpi" and x_ge_0: "0 \<le> x - k * (2 * pi)"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1643 | with lpi ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1644 | have pi_ux: "?ux \<le> pi" and ux_0: "0 \<le> real_of_float ?ux" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1645 | by simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1646 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1647 | have "(lb_cos prec ?ux) \<le> cos ?ux" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1648 | using lb_cos[OF ux_0 pi_ux] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1649 | also have "\<dots> \<le> cos (x + (-k) * (2 * pi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1650 | using cos_monotone_0_pi_le[OF x_ge_0 ux pi_ux] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1651 | by (simp only: of_int_minus | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1652 | cos_minus mult_minus_left) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1653 | finally have "(lb_cos prec ?ux) \<le> cos x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1654 | unfolding cos_periodic_int . } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1655 | note positive_ux = this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1656 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1657 | show "l \<le> cos x \<and> cos x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1658 | proof (cases "- ?lpi \<le> ?lx \<and> ?ux \<le> 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1659 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1660 | with bnds have l: "l = lb_cos prec (-?lx)" and u: "u = ub_cos prec (-?ux)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1661 | by (auto simp add: bnds_cos_def Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1662 | from True lpi[THEN le_imp_neg_le] lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1663 | have "- pi \<le> x - k * (2 * pi)" and "x - k * (2 * pi) \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1664 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1665 | with True negative_ux negative_lx show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1666 | unfolding l u by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1667 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1668 | case 1: False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1669 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1670 | proof (cases "0 \<le> ?lx \<and> ?ux \<le> ?lpi") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1671 | case True with bnds 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1672 | have l: "l = lb_cos prec ?ux" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1673 | and u: "u = ub_cos prec ?lx" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1674 | by (auto simp add: bnds_cos_def Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1675 | from True lpi lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1676 | have "0 \<le> x - k * (2 * pi)" and "x - k * (2 * pi) \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1677 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1678 | with True positive_ux positive_lx show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1679 | unfolding l u by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1680 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1681 | case 2: False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1682 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1683 | proof (cases "- ?lpi \<le> ?lx \<and> ?ux \<le> ?lpi") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1684 | case Cond: True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1685 | with bnds 1 2 have l: "l = min (lb_cos prec (-?lx)) (lb_cos prec ?ux)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1686 | and u: "u = Float 1 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1687 | by (auto simp add: bnds_cos_def Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1688 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1689 | unfolding u l using negative_lx positive_ux Cond | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1690 | by (cases "x - k * (2 * pi) < 0") (auto simp add: real_of_float_min) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1691 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1692 | case 3: False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1693 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1694 | proof (cases "0 \<le> ?lx \<and> ?ux \<le> 2 * ?lpi") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1695 | case Cond: True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1696 | with bnds 1 2 3 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1697 | have l: "l = Float (- 1) 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1698 | and u: "u = max (ub_cos prec ?lx) (ub_cos prec (- (?ux - 2 * ?lpi)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1699 | by (auto simp add: bnds_cos_def Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1700 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1701 | have "cos x \<le> real_of_float u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1702 | proof (cases "x - k * (2 * pi) < pi") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1703 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1704 | hence "x - k * (2 * pi) \<le> pi" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1705 | from positive_lx[OF Cond[THEN conjunct1] this] show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1706 | unfolding u by (simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1707 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1708 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1709 | hence "pi \<le> x - k * (2 * pi)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1710 | hence pi_x: "- pi \<le> x - k * (2 * pi) - 2 * pi" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1711 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1712 | have "?ux \<le> 2 * pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1713 | using Cond lpi by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1714 | hence "x - k * (2 * pi) - 2 * pi \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1715 | using ux by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1716 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1717 | have ux_0: "real_of_float (?ux - 2 * ?lpi) \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1718 | using Cond by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1719 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1720 | from 2 and Cond have "\<not> ?ux \<le> ?lpi" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1721 | hence "- ?lpi \<le> ?ux - 2 * ?lpi" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1722 | hence pi_ux: "- pi \<le> (?ux - 2 * ?lpi)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1723 | using lpi[THEN le_imp_neg_le] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1724 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1725 | have x_le_ux: "x - k * (2 * pi) - 2 * pi \<le> (?ux - 2 * ?lpi)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1726 | using ux lpi by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1727 | have "cos x = cos (x + (-k) * (2 * pi) + (-1::int) * (2 * pi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1728 | unfolding cos_periodic_int .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1729 | also have "\<dots> \<le> cos ((?ux - 2 * ?lpi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1730 | using cos_monotone_minus_pi_0'[OF pi_x x_le_ux ux_0] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1731 | by (simp only: minus_float.rep_eq of_int_minus of_int_1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1732 | mult_minus_left mult_1_left) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1733 | also have "\<dots> = cos ((- (?ux - 2 * ?lpi)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1734 | unfolding uminus_float.rep_eq cos_minus .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1735 | also have "\<dots> \<le> (ub_cos prec (- (?ux - 2 * ?lpi)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1736 | using lb_cos_minus[OF pi_ux ux_0] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1737 | finally show ?thesis unfolding u by (simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1738 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1739 | thus ?thesis unfolding l by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1740 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1741 | case 4: False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1742 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1743 | proof (cases "-2 * ?lpi \<le> ?lx \<and> ?ux \<le> 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1744 | case Cond: True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1745 | with bnds 1 2 3 4 have l: "l = Float (- 1) 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1746 | and u: "u = max (ub_cos prec (?lx + 2 * ?lpi)) (ub_cos prec (-?ux))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1747 | by (auto simp add: bnds_cos_def Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1748 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1749 | have "cos x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1750 | proof (cases "-pi < x - k * (2 * pi)") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1751 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1752 | hence "-pi \<le> x - k * (2 * pi)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1753 | from negative_ux[OF this Cond[THEN conjunct2]] show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1754 | unfolding u by (simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1755 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1756 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1757 | hence "x - k * (2 * pi) \<le> -pi" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1758 | hence pi_x: "x - k * (2 * pi) + 2 * pi \<le> pi" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1759 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1760 | have "-2 * pi \<le> ?lx" using Cond lpi by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1761 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1762 | hence "0 \<le> x - k * (2 * pi) + 2 * pi" using lx by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1763 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1764 | have lx_0: "0 \<le> real_of_float (?lx + 2 * ?lpi)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1765 | using Cond lpi by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1766 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1767 | from 1 and Cond have "\<not> -?lpi \<le> ?lx" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1768 | hence "?lx + 2 * ?lpi \<le> ?lpi" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1769 | hence pi_lx: "(?lx + 2 * ?lpi) \<le> pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1770 | using lpi[THEN le_imp_neg_le] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1771 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1772 | have lx_le_x: "(?lx + 2 * ?lpi) \<le> x - k * (2 * pi) + 2 * pi" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1773 | using lx lpi by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1774 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1775 | have "cos x = cos (x + (-k) * (2 * pi) + (1 :: int) * (2 * pi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1776 | unfolding cos_periodic_int .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1777 | also have "\<dots> \<le> cos ((?lx + 2 * ?lpi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1778 | using cos_monotone_0_pi_le[OF lx_0 lx_le_x pi_x] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1779 | by (simp only: minus_float.rep_eq of_int_minus of_int_1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1780 | mult_minus_left mult_1_left) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1781 | also have "\<dots> \<le> (ub_cos prec (?lx + 2 * ?lpi))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1782 | using lb_cos[OF lx_0 pi_lx] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1783 | finally show ?thesis unfolding u by (simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1784 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1785 | thus ?thesis unfolding l by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1786 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1787 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1788 | with bnds 1 2 3 4 show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1789 | by (auto simp add: bnds_cos_def Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1790 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1791 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1792 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1793 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1794 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1795 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1796 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1797 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1798 | section "Exponential function" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1799 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1800 | subsection "Compute the series of the exponential function" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1801 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1802 | fun ub_exp_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1803 | and lb_exp_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1804 | where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1805 | "ub_exp_horner prec 0 i k x = 0" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1806 | "ub_exp_horner prec (Suc n) i k x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1807 | (rapprox_rat prec 1 (int k)) (float_round_up prec (x * lb_exp_horner prec n (i + 1) (k * i) x))" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1808 | "lb_exp_horner prec 0 i k x = 0" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1809 | "lb_exp_horner prec (Suc n) i k x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1810 | (lapprox_rat prec 1 (int k)) (float_round_down prec (x * ub_exp_horner prec n (i + 1) (k * i) x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1811 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1812 | lemma bnds_exp_horner: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1813 | assumes "real_of_float x \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1814 |   shows "exp x \<in> {lb_exp_horner prec (get_even n) 1 1 x .. ub_exp_horner prec (get_odd n) 1 1 x}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1815 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1816 | have f_eq: "fact (Suc n) = fact n * ((\<lambda>i::nat. i + 1) ^^ n) 1" for n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1817 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1818 | have F: "\<And> m. ((\<lambda>i. i + 1) ^^ n) m = n + m" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1819 | by (induct n) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1820 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1821 | unfolding F by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1822 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1823 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1824 | note bounds = horner_bounds_nonpos[where f="fact" and lb="lb_exp_horner prec" and ub="ub_exp_horner prec" and j'=0 and s=1, | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1825 | OF assms f_eq lb_exp_horner.simps ub_exp_horner.simps] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1826 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1827 | have "lb_exp_horner prec (get_even n) 1 1 x \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1828 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1829 | have "lb_exp_horner prec (get_even n) 1 1 x \<le> (\<Sum>j = 0..<get_even n. 1 / (fact j) * real_of_float x ^ j)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1830 | using bounds(1) by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1831 | also have "\<dots> \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1832 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1833 | obtain t where "\<bar>t\<bar> \<le> \<bar>real_of_float x\<bar>" and "exp x = (\<Sum>m = 0..<get_even n. real_of_float x ^ m / (fact m)) + exp t / (fact (get_even n)) * (real_of_float x) ^ (get_even n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1834 | using Maclaurin_exp_le unfolding atLeast0LessThan by blast | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1835 | moreover have "0 \<le> exp t / (fact (get_even n)) * (real_of_float x) ^ (get_even n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1836 | by (auto simp: zero_le_even_power) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1837 | ultimately show ?thesis using get_odd exp_gt_zero by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1838 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1839 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1840 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1841 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1842 | have "exp x \<le> ub_exp_horner prec (get_odd n) 1 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1843 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1844 | have x_less_zero: "real_of_float x ^ get_odd n \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1845 | proof (cases "real_of_float x = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1846 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1847 | have "(get_odd n) \<noteq> 0" using get_odd[THEN odd_pos] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1848 | thus ?thesis unfolding True power_0_left by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1849 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1850 | case False hence "real_of_float x < 0" using \<open>real_of_float x \<le> 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1851 | show ?thesis by (rule less_imp_le, auto simp add: \<open>real_of_float x < 0\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1852 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1853 | obtain t where "\<bar>t\<bar> \<le> \<bar>real_of_float x\<bar>" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1854 | and "exp x = (\<Sum>m = 0..<get_odd n. (real_of_float x) ^ m / (fact m)) + exp t / (fact (get_odd n)) * (real_of_float x) ^ (get_odd n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1855 | using Maclaurin_exp_le unfolding atLeast0LessThan by blast | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1856 | moreover have "exp t / (fact (get_odd n)) * (real_of_float x) ^ (get_odd n) \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1857 | by (auto intro!: mult_nonneg_nonpos divide_nonpos_pos simp add: x_less_zero) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1858 | ultimately have "exp x \<le> (\<Sum>j = 0..<get_odd n. 1 / (fact j) * real_of_float x ^ j)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1859 | using get_odd exp_gt_zero by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1860 | also have "\<dots> \<le> ub_exp_horner prec (get_odd n) 1 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1861 | using bounds(2) by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1862 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1863 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1864 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1865 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1866 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1867 | lemma ub_exp_horner_nonneg: "real_of_float x \<le> 0 \<Longrightarrow> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1868 | 0 \<le> real_of_float (ub_exp_horner prec (get_odd n) (Suc 0) (Suc 0) x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1869 | using bnds_exp_horner[of x prec n] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1870 | by (intro order_trans[OF exp_ge_zero]) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1871 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1872 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1873 | subsection "Compute the exponential function on the entire domain" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1874 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1875 | function ub_exp :: "nat \<Rightarrow> float \<Rightarrow> float" and lb_exp :: "nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1876 | "lb_exp prec x = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1877 | (if 0 < x then float_divl prec 1 (ub_exp prec (-x)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1878 | else | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1879 | let | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1880 | horner = (\<lambda> x. let y = lb_exp_horner prec (get_even (prec + 2)) 1 1 x in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1881 | if y \<le> 0 then Float 1 (- 2) else y) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1882 | in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1883 | if x < - 1 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1884 | power_down_fl prec (horner (float_divl prec x (- floor_fl x))) (nat (- int_floor_fl x)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1885 | else horner x)" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1886 | "ub_exp prec x = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1887 | (if 0 < x then float_divr prec 1 (lb_exp prec (-x)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1888 | else if x < - 1 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1889 | power_up_fl prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1890 | (ub_exp_horner prec (get_odd (prec + 2)) 1 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1891 | (float_divr prec x (- floor_fl x))) (nat (- int_floor_fl x)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1892 | else ub_exp_horner prec (get_odd (prec + 2)) 1 1 x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1893 | by pat_completeness auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1894 | termination | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1895 | by (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if 0 < x then 1 else 0))") auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1896 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1897 | lemma exp_m1_ge_quarter: "(1 / 4 :: real) \<le> exp (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1898 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1899 | have eq4: "4 = Suc (Suc (Suc (Suc 0)))" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1900 | have "1 / 4 = (Float 1 (- 2))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1901 | unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1902 | also have "\<dots> \<le> lb_exp_horner 3 (get_even 3) 1 1 (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1903 | by (subst less_eq_float.rep_eq [symmetric]) code_simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1904 | also have "\<dots> \<le> exp (- 1 :: float)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1905 | using bnds_exp_horner[where x="- 1"] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1906 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1907 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1908 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1909 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1910 | lemma lb_exp_pos: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1911 | assumes "\<not> 0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1912 | shows "0 < lb_exp prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1913 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1914 | let "?lb_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1915 | let "?horner x" = "let y = ?lb_horner x in if y \<le> 0 then Float 1 (- 2) else y" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1916 | have pos_horner: "0 < ?horner x" for x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1917 | unfolding Let_def by (cases "?lb_horner x \<le> 0") auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1918 | moreover have "0 < real_of_float ((?horner x) ^ num)" for x :: float and num :: nat | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1919 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1920 | have "0 < real_of_float (?horner x) ^ num" using \<open>0 < ?horner x\<close> by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1921 | also have "\<dots> = (?horner x) ^ num" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1922 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1923 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1924 | ultimately show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1925 | unfolding lb_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1926 | by (cases "floor_fl x", cases "x < - 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1927 | (auto simp: real_power_up_fl real_power_down_fl intro!: power_up_less power_down_pos) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1928 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1929 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1930 | lemma exp_boundaries': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1931 | assumes "x \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1932 |   shows "exp x \<in> { (lb_exp prec x) .. (ub_exp prec x)}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1933 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1934 | let "?lb_exp_horner x" = "lb_exp_horner prec (get_even (prec + 2)) 1 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1935 | let "?ub_exp_horner x" = "ub_exp_horner prec (get_odd (prec + 2)) 1 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1936 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1937 | have "real_of_float x \<le> 0" and "\<not> x > 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1938 | using \<open>x \<le> 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1939 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1940 | proof (cases "x < - 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1941 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1942 | hence "- 1 \<le> real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1943 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1944 | proof (cases "?lb_exp_horner x \<le> 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1945 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1946 | from \<open>\<not> x < - 1\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1947 | have "- 1 \<le> real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1948 | hence "exp (- 1) \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1949 | unfolding exp_le_cancel_iff . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1950 | from order_trans[OF exp_m1_ge_quarter this] have "Float 1 (- 2) \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1951 | unfolding Float_num . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1952 | with True show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1953 | using bnds_exp_horner \<open>real_of_float x \<le> 0\<close> \<open>\<not> x > 0\<close> \<open>\<not> x < - 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1954 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1955 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1956 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1957 | using bnds_exp_horner \<open>real_of_float x \<le> 0\<close> \<open>\<not> x > 0\<close> \<open>\<not> x < - 1\<close> by (auto simp add: Let_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1958 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1959 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1960 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1961 | let ?num = "nat (- int_floor_fl x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1962 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1963 | have "real_of_int (int_floor_fl x) < - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1964 | using int_floor_fl[of x] \<open>x < - 1\<close> by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1965 | hence "real_of_int (int_floor_fl x) < 0" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1966 | hence "int_floor_fl x < 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1967 | hence "1 \<le> - int_floor_fl x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1968 | hence "0 < nat (- int_floor_fl x)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1969 | hence "0 < ?num" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1970 | hence "real ?num \<noteq> 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1971 | have num_eq: "real ?num = - int_floor_fl x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1972 | using \<open>0 < nat (- int_floor_fl x)\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1973 | have "0 < - int_floor_fl x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1974 | using \<open>0 < ?num\<close>[unfolded of_nat_less_iff[symmetric]] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1975 | hence "real_of_int (int_floor_fl x) < 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1976 | unfolding less_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1977 | have fl_eq: "real_of_int (- int_floor_fl x) = real_of_float (- floor_fl x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1978 | by (simp add: floor_fl_def int_floor_fl_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1979 | from \<open>0 < - int_floor_fl x\<close> have "0 \<le> real_of_float (- floor_fl x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1980 | by (simp add: floor_fl_def int_floor_fl_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1981 | from \<open>real_of_int (int_floor_fl x) < 0\<close> have "real_of_float (floor_fl x) < 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1982 | by (simp add: floor_fl_def int_floor_fl_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1983 | have "exp x \<le> ub_exp prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1984 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1985 | have div_less_zero: "real_of_float (float_divr prec x (- floor_fl x)) \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1986 | using float_divr_nonpos_pos_upper_bound[OF \<open>real_of_float x \<le> 0\<close> \<open>0 \<le> real_of_float (- floor_fl x)\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1987 | unfolding less_eq_float_def zero_float.rep_eq . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1988 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1989 | have "exp x = exp (?num * (x / ?num))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1990 | using \<open>real ?num \<noteq> 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1991 | also have "\<dots> = exp (x / ?num) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1992 | unfolding exp_of_nat_mult .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1993 | also have "\<dots> \<le> exp (float_divr prec x (- floor_fl x)) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1994 | unfolding num_eq fl_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1995 | by (rule power_mono, rule exp_le_cancel_iff[THEN iffD2], rule float_divr) auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1996 | also have "\<dots> \<le> (?ub_exp_horner (float_divr prec x (- floor_fl x))) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1997 | unfolding real_of_float_power | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1998 | by (rule power_mono, rule bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct2], auto) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 1999 | also have "\<dots> \<le> real_of_float (power_up_fl prec (?ub_exp_horner (float_divr prec x (- floor_fl x))) ?num)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2000 | by (auto simp add: real_power_up_fl intro!: power_up ub_exp_horner_nonneg div_less_zero) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2001 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2002 | unfolding ub_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] if_P[OF \<open>x < - 1\<close>] floor_fl_def Let_def . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2003 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2004 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2005 | have "lb_exp prec x \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2006 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2007 | let ?divl = "float_divl prec x (- floor_fl x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2008 | let ?horner = "?lb_exp_horner ?divl" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2009 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2010 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2011 | proof (cases "?horner \<le> 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2012 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2013 | hence "0 \<le> real_of_float ?horner" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2014 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2015 | have div_less_zero: "real_of_float (float_divl prec x (- floor_fl x)) \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2016 | using \<open>real_of_float (floor_fl x) < 0\<close> \<open>real_of_float x \<le> 0\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2017 | by (auto intro!: order_trans[OF float_divl] divide_nonpos_neg) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2018 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2019 | have "(?lb_exp_horner (float_divl prec x (- floor_fl x))) ^ ?num \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2020 | exp (float_divl prec x (- floor_fl x)) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2021 | using \<open>0 \<le> real_of_float ?horner\<close>[unfolded floor_fl_def[symmetric]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2022 | bnds_exp_horner[OF div_less_zero, unfolded atLeastAtMost_iff, THEN conjunct1] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2023 | by (auto intro!: power_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2024 | also have "\<dots> \<le> exp (x / ?num) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2025 | unfolding num_eq fl_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2026 | using float_divl by (auto intro!: power_mono simp del: uminus_float.rep_eq) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2027 | also have "\<dots> = exp (?num * (x / ?num))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2028 | unfolding exp_of_nat_mult .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2029 | also have "\<dots> = exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2030 | using \<open>real ?num \<noteq> 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2031 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2032 | using False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2033 | unfolding lb_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] if_P[OF \<open>x < - 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2034 | int_floor_fl_def Let_def if_not_P[OF False] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2035 | by (auto simp: real_power_down_fl intro!: power_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2036 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2037 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2038 | have "power_down_fl prec (Float 1 (- 2)) ?num \<le> (Float 1 (- 2)) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2039 | by (metis Float_le_zero_iff less_imp_le linorder_not_less | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2040 | not_numeral_le_zero numeral_One power_down_fl) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2041 | then have "power_down_fl prec (Float 1 (- 2)) ?num \<le> real_of_float (Float 1 (- 2)) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2042 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2043 | also | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2044 | have "real_of_float (floor_fl x) \<noteq> 0" and "real_of_float (floor_fl x) \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2045 | using \<open>real_of_float (floor_fl x) < 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2046 | from divide_right_mono_neg[OF floor_fl[of x] \<open>real_of_float (floor_fl x) \<le> 0\<close>, unfolded divide_self[OF \<open>real_of_float (floor_fl x) \<noteq> 0\<close>]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2047 | have "- 1 \<le> x / (- floor_fl x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2048 | unfolding minus_float.rep_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2049 | from order_trans[OF exp_m1_ge_quarter this[unfolded exp_le_cancel_iff[where x="- 1", symmetric]]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2050 | have "Float 1 (- 2) \<le> exp (x / (- floor_fl x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2051 | unfolding Float_num . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2052 | hence "real_of_float (Float 1 (- 2)) ^ ?num \<le> exp (x / (- floor_fl x)) ^ ?num" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2053 | by (metis Float_num(5) power_mono zero_le_divide_1_iff zero_le_numeral) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2054 | also have "\<dots> = exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2055 | unfolding num_eq fl_eq exp_of_nat_mult[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2056 | using \<open>real_of_float (floor_fl x) \<noteq> 0\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2057 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2058 | unfolding lb_exp.simps if_not_P[OF \<open>\<not> 0 < x\<close>] if_P[OF \<open>x < - 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2059 | int_floor_fl_def Let_def if_P[OF True] real_of_float_power . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2060 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2061 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2062 | ultimately show ?thesis by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2063 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2064 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2065 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2066 | lemma exp_boundaries: "exp x \<in> { lb_exp prec x .. ub_exp prec x }"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2067 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2068 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2069 | proof (cases "0 < x") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2070 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2071 | hence "x \<le> 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2072 | from exp_boundaries'[OF this] show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2073 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2074 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2075 | hence "-x \<le> 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2076 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2077 | have "lb_exp prec x \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2078 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2079 | from exp_boundaries'[OF \<open>-x \<le> 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2080 | have ub_exp: "exp (- real_of_float x) \<le> ub_exp prec (-x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2081 | unfolding atLeastAtMost_iff minus_float.rep_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2082 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2083 | have "float_divl prec 1 (ub_exp prec (-x)) \<le> 1 / ub_exp prec (-x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2084 | using float_divl[where x=1] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2085 | also have "\<dots> \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2086 | using ub_exp[unfolded inverse_le_iff_le[OF order_less_le_trans[OF exp_gt_zero ub_exp] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2087 | exp_gt_zero, symmetric]] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2088 | unfolding exp_minus nonzero_inverse_inverse_eq[OF exp_not_eq_zero] inverse_eq_divide | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2089 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2090 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2091 | unfolding lb_exp.simps if_P[OF True] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2092 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2093 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2094 | have "exp x \<le> ub_exp prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2095 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2096 | have "\<not> 0 < -x" using \<open>0 < x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2097 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2098 | from exp_boundaries'[OF \<open>-x \<le> 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2099 | have lb_exp: "lb_exp prec (-x) \<le> exp (- real_of_float x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2100 | unfolding atLeastAtMost_iff minus_float.rep_eq by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2101 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2102 | have "exp x \<le> (1 :: float) / lb_exp prec (-x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2103 | using lb_exp lb_exp_pos[OF \<open>\<not> 0 < -x\<close>, of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2104 | by (simp del: lb_exp.simps add: exp_minus field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2105 | also have "\<dots> \<le> float_divr prec 1 (lb_exp prec (-x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2106 | using float_divr . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2107 | finally show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2108 | unfolding ub_exp.simps if_P[OF True] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2109 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2110 | ultimately show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2111 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2112 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2113 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2114 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2115 | lemma bnds_exp: "\<forall>(x::real) lx ux. (l, u) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2116 |   (lb_exp prec lx, ub_exp prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> exp x \<and> exp x \<le> u"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2117 | proof (rule allI, rule allI, rule allI, rule impI) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2118 | fix x :: real and lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2119 |   assume "(l, u) = (lb_exp prec lx, ub_exp prec ux) \<and> x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2120 |   hence l: "lb_exp prec lx = l " and u: "ub_exp prec ux = u" and x: "x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2121 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2122 | show "l \<le> exp x \<and> exp x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2123 | proof | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2124 | show "l \<le> exp x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2125 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2126 | from exp_boundaries[of lx prec, unfolded l] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2127 | have "l \<le> exp lx" by (auto simp del: lb_exp.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2128 | also have "\<dots> \<le> exp x" using x by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2129 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2130 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2131 | show "exp x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2132 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2133 | have "exp x \<le> exp ux" using x by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2134 | also have "\<dots> \<le> u" using exp_boundaries[of ux prec, unfolded u] by (auto simp del: ub_exp.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2135 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2136 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2137 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2138 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2139 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2140 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2141 | section "Logarithm" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2142 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2143 | subsection "Compute the logarithm series" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2144 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2145 | fun ub_ln_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2146 | and lb_ln_horner :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> float \<Rightarrow> float" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2147 | "ub_ln_horner prec 0 i x = 0" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2148 | "ub_ln_horner prec (Suc n) i x = float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2149 | (rapprox_rat prec 1 (int i)) (- float_round_down prec (x * lb_ln_horner prec n (Suc i) x))" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2150 | "lb_ln_horner prec 0 i x = 0" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2151 | "lb_ln_horner prec (Suc n) i x = float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2152 | (lapprox_rat prec 1 (int i)) (- float_round_up prec (x * ub_ln_horner prec n (Suc i) x))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2153 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2154 | lemma ln_bounds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2155 | assumes "0 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2156 | and "x < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2157 | shows "(\<Sum>i=0..<2*n. (- 1) ^ i * (1 / real (i + 1)) * x ^ (Suc i)) \<le> ln (x + 1)" (is "?lb") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2158 | and "ln (x + 1) \<le> (\<Sum>i=0..<2*n + 1. (- 1) ^ i * (1 / real (i + 1)) * x ^ (Suc i))" (is "?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2159 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2160 | let "?a n" = "(1/real (n +1)) * x ^ (Suc n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2161 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2162 | have ln_eq: "(\<Sum> i. (- 1) ^ i * ?a i) = ln (x + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2163 | using ln_series[of "x + 1"] \<open>0 \<le> x\<close> \<open>x < 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2164 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2165 | have "norm x < 1" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2166 | have "?a \<longlonglongrightarrow> 0" unfolding Suc_eq_plus1[symmetric] inverse_eq_divide[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2167 | using tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_Suc[OF LIMSEQ_power_zero[OF \<open>norm x < 1\<close>]]] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2168 | have "0 \<le> ?a n" for n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2169 | by (rule mult_nonneg_nonneg) (auto simp: \<open>0 \<le> x\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2170 | have "?a (Suc n) \<le> ?a n" for n | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2171 | unfolding inverse_eq_divide[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2172 | proof (rule mult_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2173 | show "0 \<le> x ^ Suc (Suc n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2174 | by (auto simp add: \<open>0 \<le> x\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2175 | have "x ^ Suc (Suc n) \<le> x ^ Suc n * 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2176 | unfolding power_Suc2 mult.assoc[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2177 | by (rule mult_left_mono, fact less_imp_le[OF \<open>x < 1\<close>]) (auto simp: \<open>0 \<le> x\<close>) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2178 | thus "x ^ Suc (Suc n) \<le> x ^ Suc n" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2179 | qed auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2180 | from summable_Leibniz'(2,4)[OF \<open>?a \<longlonglongrightarrow> 0\<close> \<open>\<And>n. 0 \<le> ?a n\<close>, OF \<open>\<And>n. ?a (Suc n) \<le> ?a n\<close>, unfolded ln_eq] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2181 | show ?lb and ?ub | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2182 | unfolding atLeast0LessThan by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2183 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2184 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2185 | lemma ln_float_bounds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2186 | assumes "0 \<le> real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2187 | and "real_of_float x < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2188 | shows "x * lb_ln_horner prec (get_even n) 1 x \<le> ln (x + 1)" (is "?lb \<le> ?ln") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2189 | and "ln (x + 1) \<le> x * ub_ln_horner prec (get_odd n) 1 x" (is "?ln \<le> ?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2190 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2191 | obtain ev where ev: "get_even n = 2 * ev" using get_even_double .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2192 | obtain od where od: "get_odd n = 2 * od + 1" using get_odd_double .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2193 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2194 | let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real_of_float x)^(Suc n)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2195 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2196 |   have "?lb \<le> sum ?s {0 ..< 2 * ev}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2197 | unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2198 | unfolding mult.commute[of "real_of_float x"] ev | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2199 | using horner_bounds(1)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2200 | and lb="\<lambda>n i k x. lb_ln_horner prec n k x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2201 | and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*ev", | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2202 | OF \<open>0 \<le> real_of_float x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2203 | unfolding real_of_float_power | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2204 | by (rule mult_right_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2205 | also have "\<dots> \<le> ?ln" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2206 | using ln_bounds(1)[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x < 1\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2207 | finally show "?lb \<le> ?ln" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2208 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2209 |   have "?ln \<le> sum ?s {0 ..< 2 * od + 1}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2210 | using ln_bounds(2)[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x < 1\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2211 | also have "\<dots> \<le> ?ub" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2212 | unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2213 | unfolding mult.commute[of "real_of_float x"] od | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2214 | using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1", | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2215 | OF \<open>0 \<le> real_of_float x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real_of_float x\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2216 | unfolding real_of_float_power | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2217 | by (rule mult_right_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2218 | finally show "?ln \<le> ?ub" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2219 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2220 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2221 | lemma ln_add: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2222 | fixes x :: real | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2223 | assumes "0 < x" and "0 < y" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2224 | shows "ln (x + y) = ln x + ln (1 + y / x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2225 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2226 | have "x \<noteq> 0" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2227 | have "x + y = x * (1 + y / x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2228 | unfolding distrib_left times_divide_eq_right nonzero_mult_div_cancel_left[OF \<open>x \<noteq> 0\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2229 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2230 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2231 | have "0 < y / x" using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2232 | hence "0 < 1 + y / x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2233 | ultimately show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2234 | using ln_mult assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2235 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2236 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2237 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2238 | subsection "Compute the logarithm of 2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2239 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2240 | definition ub_ln2 where "ub_ln2 prec = (let third = rapprox_rat (max prec 1) 1 3 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2241 | in float_plus_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2242 | ((Float 1 (- 1) * ub_ln_horner prec (get_odd prec) 1 (Float 1 (- 1)))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2243 | (float_round_up prec (third * ub_ln_horner prec (get_odd prec) 1 third)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2244 | definition lb_ln2 where "lb_ln2 prec = (let third = lapprox_rat prec 1 3 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2245 | in float_plus_down prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2246 | ((Float 1 (- 1) * lb_ln_horner prec (get_even prec) 1 (Float 1 (- 1)))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2247 | (float_round_down prec (third * lb_ln_horner prec (get_even prec) 1 third)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2248 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2249 | lemma ub_ln2: "ln 2 \<le> ub_ln2 prec" (is "?ub_ln2") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2250 | and lb_ln2: "lb_ln2 prec \<le> ln 2" (is "?lb_ln2") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2251 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2252 | let ?uthird = "rapprox_rat (max prec 1) 1 3" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2253 | let ?lthird = "lapprox_rat prec 1 3" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2254 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2255 | have ln2_sum: "ln 2 = ln (1/2 + 1) + ln (1 / 3 + 1::real)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2256 | using ln_add[of "3 / 2" "1 / 2"] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2257 | have lb3: "?lthird \<le> 1 / 3" using lapprox_rat[of prec 1 3] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2258 | hence lb3_ub: "real_of_float ?lthird < 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2259 | have lb3_lb: "0 \<le> real_of_float ?lthird" using lapprox_rat_nonneg[of 1 3] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2260 | have ub3: "1 / 3 \<le> ?uthird" using rapprox_rat[of 1 3] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2261 | hence ub3_lb: "0 \<le> real_of_float ?uthird" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2262 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2263 | have lb2: "0 \<le> real_of_float (Float 1 (- 1))" and ub2: "real_of_float (Float 1 (- 1)) < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2264 | unfolding Float_num by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2265 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2266 | have "0 \<le> (1::int)" and "0 < (3::int)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2267 | have ub3_ub: "real_of_float ?uthird < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2268 | by (simp add: Float.compute_rapprox_rat Float.compute_lapprox_rat rapprox_posrat_less1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2269 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2270 | have third_gt0: "(0 :: real) < 1 / 3 + 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2271 | have uthird_gt0: "0 < real_of_float ?uthird + 1" using ub3_lb by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2272 | have lthird_gt0: "0 < real_of_float ?lthird + 1" using lb3_lb by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2273 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2274 | show ?ub_ln2 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2275 | unfolding ub_ln2_def Let_def ln2_sum Float_num(4)[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2276 | proof (rule float_plus_up_le, rule add_mono, fact ln_float_bounds(2)[OF lb2 ub2]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2277 | have "ln (1 / 3 + 1) \<le> ln (real_of_float ?uthird + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2278 | unfolding ln_le_cancel_iff[OF third_gt0 uthird_gt0] using ub3 by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2279 | also have "\<dots> \<le> ?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2280 | using ln_float_bounds(2)[OF ub3_lb ub3_ub] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2281 | also note float_round_up | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2282 | finally show "ln (1 / 3 + 1) \<le> float_round_up prec (?uthird * ub_ln_horner prec (get_odd prec) 1 ?uthird)" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2283 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2284 | show ?lb_ln2 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2285 | unfolding lb_ln2_def Let_def ln2_sum Float_num(4)[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2286 | proof (rule float_plus_down_le, rule add_mono, fact ln_float_bounds(1)[OF lb2 ub2]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2287 | have "?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird \<le> ln (real_of_float ?lthird + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2288 | using ln_float_bounds(1)[OF lb3_lb lb3_ub] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2289 | note float_round_down_le[OF this] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2290 | also have "\<dots> \<le> ln (1 / 3 + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2291 | unfolding ln_le_cancel_iff[OF lthird_gt0 third_gt0] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2292 | using lb3 by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2293 | finally show "float_round_down prec (?lthird * lb_ln_horner prec (get_even prec) 1 ?lthird) \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2294 | ln (1 / 3 + 1)" . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2295 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2296 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2297 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2298 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2299 | subsection "Compute the logarithm in the entire domain" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2300 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2301 | function ub_ln :: "nat \<Rightarrow> float \<Rightarrow> float option" and lb_ln :: "nat \<Rightarrow> float \<Rightarrow> float option" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2302 | "ub_ln prec x = (if x \<le> 0 then None | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2303 | else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2304 | else let horner = \<lambda>x. float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x) in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2305 | if x \<le> Float 3 (- 1) then Some (horner (x - 1)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2306 | else if x < Float 1 1 then Some (float_round_up prec (horner (Float 1 (- 1)) + horner (x * rapprox_rat prec 2 3 - 1))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2307 | else let l = bitlen (mantissa x) - 1 in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2308 | Some (float_plus_up prec (float_round_up prec (ub_ln2 prec * (Float (exponent x + l) 0))) (horner (Float (mantissa x) (- l) - 1))))" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2309 | "lb_ln prec x = (if x \<le> 0 then None | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2310 | else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2311 | else let horner = \<lambda>x. float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x) in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2312 | if x \<le> Float 3 (- 1) then Some (horner (x - 1)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2313 | else if x < Float 1 1 then Some (float_round_down prec (horner (Float 1 (- 1)) + | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2314 | horner (max (x * lapprox_rat prec 2 3 - 1) 0))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2315 | else let l = bitlen (mantissa x) - 1 in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2316 | Some (float_plus_down prec (float_round_down prec (lb_ln2 prec * (Float (exponent x + l) 0))) (horner (Float (mantissa x) (- l) - 1))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2317 | by pat_completeness auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2318 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2319 | termination | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2320 | proof (relation "measure (\<lambda> v. let (prec, x) = case_sum id id v in (if x < 1 then 1 else 0))", auto) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2321 | fix prec and x :: float | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2322 | assume "\<not> real_of_float x \<le> 0" and "real_of_float x < 1" and "real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2323 | hence "0 < real_of_float x" "1 \<le> max prec (Suc 0)" "real_of_float x < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2324 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2325 | from float_divl_pos_less1_bound[OF \<open>0 < real_of_float x\<close> \<open>real_of_float x < 1\<close>[THEN less_imp_le] \<open>1 \<le> max prec (Suc 0)\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2326 | show False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2327 | using \<open>real_of_float (float_divl (max prec (Suc 0)) 1 x) < 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2328 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2329 | fix prec x | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2330 | assume "\<not> real_of_float x \<le> 0" and "real_of_float x < 1" and "real_of_float (float_divr prec 1 x) < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2331 | hence "0 < x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2332 | from float_divr_pos_less1_lower_bound[OF \<open>0 < x\<close>, of prec] \<open>real_of_float x < 1\<close> show False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2333 | using \<open>real_of_float (float_divr prec 1 x) < 1\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2334 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2335 | |
| 67573 | 2336 | lemmas float_pos_eq_mantissa_pos = mantissa_pos_iff[symmetric] | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2337 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2338 | lemma Float_pos_eq_mantissa_pos: "Float m e > 0 \<longleftrightarrow> m > 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2339 | using powr_gt_zero[of 2 "e"] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2340 | by (auto simp add: zero_less_mult_iff zero_float_def simp del: powr_gt_zero dest: less_zeroE) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2341 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2342 | lemma Float_representation_aux: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2343 | fixes m e | 
| 67573 | 2344 | defines [THEN meta_eq_to_obj_eq]: "x \<equiv> Float m e" | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2345 | assumes "x > 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2346 | shows "Float (exponent x + (bitlen (mantissa x) - 1)) 0 = Float (e + (bitlen m - 1)) 0" (is ?th1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2347 | and "Float (mantissa x) (- (bitlen (mantissa x) - 1)) = Float m ( - (bitlen m - 1))" (is ?th2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2348 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2349 | from assms have mantissa_pos: "m > 0" "mantissa x > 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2350 | using Float_pos_eq_mantissa_pos[of m e] float_pos_eq_mantissa_pos[of x] by simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2351 | thus ?th1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2352 | using bitlen_Float[of m e] assms | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2353 | by (auto simp add: zero_less_mult_iff intro!: arg_cong2[where f=Float]) | 
| 67573 | 2354 | have "x \<noteq> 0" | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2355 | unfolding zero_float_def[symmetric] using \<open>0 < x\<close> by auto | 
| 67573 | 2356 | from denormalize_shift[OF x_def this] obtain i where | 
| 2357 | i: "m = mantissa x * 2 ^ i" "e = exponent x - int i" . | |
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2358 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2359 | have "2 powr (1 - (real_of_int (bitlen (mantissa x)) + real_of_int i)) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2360 | 2 powr (1 - (real_of_int (bitlen (mantissa x)))) * inverse (2 powr (real i))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2361 | by (simp add: powr_minus[symmetric] powr_add[symmetric] field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2362 | hence "real_of_int (mantissa x) * 2 powr (1 - real_of_int (bitlen (mantissa x))) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2363 | (real_of_int (mantissa x) * 2 ^ i) * 2 powr (1 - real_of_int (bitlen (mantissa x * 2 ^ i)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2364 | using \<open>mantissa x > 0\<close> by (simp add: powr_realpow) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2365 | then show ?th2 | 
| 67573 | 2366 | unfolding i | 
| 2367 | by (auto simp: real_of_float_eq) | |
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2368 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2369 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2370 | lemma compute_ln[code]: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2371 | fixes m e | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2372 | defines "x \<equiv> Float m e" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2373 | shows "ub_ln prec x = (if x \<le> 0 then None | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2374 | else if x < 1 then Some (- the (lb_ln prec (float_divl (max prec 1) 1 x))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2375 | else let horner = \<lambda>x. float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x) in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2376 | if x \<le> Float 3 (- 1) then Some (horner (x - 1)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2377 | else if x < Float 1 1 then Some (float_round_up prec (horner (Float 1 (- 1)) + horner (x * rapprox_rat prec 2 3 - 1))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2378 | else let l = bitlen m - 1 in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2379 | Some (float_plus_up prec (float_round_up prec (ub_ln2 prec * (Float (e + l) 0))) (horner (Float m (- l) - 1))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2380 | (is ?th1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2381 | and "lb_ln prec x = (if x \<le> 0 then None | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2382 | else if x < 1 then Some (- the (ub_ln prec (float_divr prec 1 x))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2383 | else let horner = \<lambda>x. float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x) in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2384 | if x \<le> Float 3 (- 1) then Some (horner (x - 1)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2385 | else if x < Float 1 1 then Some (float_round_down prec (horner (Float 1 (- 1)) + | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2386 | horner (max (x * lapprox_rat prec 2 3 - 1) 0))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2387 | else let l = bitlen m - 1 in | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2388 | Some (float_plus_down prec (float_round_down prec (lb_ln2 prec * (Float (e + l) 0))) (horner (Float m (- l) - 1))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2389 | (is ?th2) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2390 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2391 | from assms Float_pos_eq_mantissa_pos have "x > 0 \<Longrightarrow> m > 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2392 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2393 | thus ?th1 ?th2 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2394 | using Float_representation_aux[of m e] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2395 | unfolding x_def[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2396 | by (auto dest: not_le_imp_less) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2397 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2398 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2399 | lemma ln_shifted_float: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2400 | assumes "0 < m" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2401 | shows "ln (Float m e) = ln 2 * (e + (bitlen m - 1)) + ln (Float m (- (bitlen m - 1)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2402 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2403 | let ?B = "2^nat (bitlen m - 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2404 | define bl where "bl = bitlen m - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2405 | have "0 < real_of_int m" and "\<And>X. (0 :: real) < 2^X" and "0 < (2 :: real)" and "m \<noteq> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2406 | using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2407 | hence "0 \<le> bl" by (simp add: bitlen_alt_def bl_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2408 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2409 | proof (cases "0 \<le> e") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2410 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2411 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2412 | unfolding bl_def[symmetric] using \<open>0 < real_of_int m\<close> \<open>0 \<le> bl\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2413 | apply (simp add: ln_mult) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2414 | apply (cases "e=0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2415 | apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2416 | apply (cases "bl = 0", simp_all add: powr_minus ln_inverse ln_powr field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2417 | done | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2418 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2419 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2420 | hence "0 < -e" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2421 | have lne: "ln (2 powr real_of_int e) = ln (inverse (2 powr - e))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2422 | by (simp add: powr_minus) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2423 | hence pow_gt0: "(0::real) < 2^nat (-e)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2424 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2425 | hence inv_gt0: "(0::real) < inverse (2^nat (-e))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2426 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2427 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2428 | using False unfolding bl_def[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2429 | using \<open>0 < real_of_int m\<close> \<open>0 \<le> bl\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2430 | by (auto simp add: lne ln_mult ln_powr ln_div field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2431 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2432 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2433 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2434 | lemma ub_ln_lb_ln_bounds': | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2435 | assumes "1 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2436 | shows "the (lb_ln prec x) \<le> ln x \<and> ln x \<le> the (ub_ln prec x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2437 | (is "?lb \<le> ?ln \<and> ?ln \<le> ?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2438 | proof (cases "x < Float 1 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2439 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2440 | hence "real_of_float (x - 1) < 1" and "real_of_float x < 2" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2441 | have "\<not> x \<le> 0" and "\<not> x < 1" using \<open>1 \<le> x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2442 | hence "0 \<le> real_of_float (x - 1)" using \<open>1 \<le> x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2443 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2444 | have [simp]: "(Float 3 (- 1)) = 3 / 2" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2445 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2446 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2447 | proof (cases "x \<le> Float 3 (- 1)") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2448 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2449 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2450 | unfolding lb_ln.simps | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2451 | unfolding ub_ln.simps Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2452 | using ln_float_bounds[OF \<open>0 \<le> real_of_float (x - 1)\<close> \<open>real_of_float (x - 1) < 1\<close>, of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2453 | \<open>\<not> x \<le> 0\<close> \<open>\<not> x < 1\<close> True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2454 | by (auto intro!: float_round_down_le float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2455 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2456 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2457 | hence *: "3 / 2 < x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2458 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2459 | with ln_add[of "3 / 2" "x - 3 / 2"] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2460 | have add: "ln x = ln (3 / 2) + ln (real_of_float x * 2 / 3)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2461 | by (auto simp add: algebra_simps diff_divide_distrib) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2462 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2463 | let "?ub_horner x" = "float_round_up prec (x * ub_ln_horner prec (get_odd prec) 1 x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2464 | let "?lb_horner x" = "float_round_down prec (x * lb_ln_horner prec (get_even prec) 1 x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2465 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2466 |     { have up: "real_of_float (rapprox_rat prec 2 3) \<le> 1"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2467 | by (rule rapprox_rat_le1) simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2468 | have low: "2 / 3 \<le> rapprox_rat prec 2 3" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2469 | by (rule order_trans[OF _ rapprox_rat]) simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2470 | from mult_less_le_imp_less[OF * low] * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2471 | have pos: "0 < real_of_float (x * rapprox_rat prec 2 3 - 1)" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2472 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2473 | have "ln (real_of_float x * 2/3) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2474 | \<le> ln (real_of_float (x * rapprox_rat prec 2 3 - 1) + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2475 | proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2476 | show "real_of_float x * 2 / 3 \<le> real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2477 | using * low by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2478 | show "0 < real_of_float x * 2 / 3" using * by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2479 | show "0 < real_of_float (x * rapprox_rat prec 2 3 - 1) + 1" using pos by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2480 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2481 | also have "\<dots> \<le> ?ub_horner (x * rapprox_rat prec 2 3 - 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2482 | proof (rule float_round_up_le, rule ln_float_bounds(2)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2483 | from mult_less_le_imp_less[OF \<open>real_of_float x < 2\<close> up] low * | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2484 | show "real_of_float (x * rapprox_rat prec 2 3 - 1) < 1" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2485 | show "0 \<le> real_of_float (x * rapprox_rat prec 2 3 - 1)" using pos by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2486 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2487 | finally have "ln x \<le> ?ub_horner (Float 1 (-1)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2488 | + ?ub_horner ((x * rapprox_rat prec 2 3 - 1))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2489 | using ln_float_bounds(2)[of "Float 1 (- 1)" prec prec] add | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2490 | by (auto intro!: add_mono float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2491 | note float_round_up_le[OF this, of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2492 | } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2493 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2494 |     { let ?max = "max (x * lapprox_rat prec 2 3 - 1) 0"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2495 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2496 | have up: "lapprox_rat prec 2 3 \<le> 2/3" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2497 | by (rule order_trans[OF lapprox_rat], simp) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2498 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2499 | have low: "0 \<le> real_of_float (lapprox_rat prec 2 3)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2500 | using lapprox_rat_nonneg[of 2 3 prec] by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2501 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2502 | have "?lb_horner ?max | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2503 | \<le> ln (real_of_float ?max + 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2504 | proof (rule float_round_down_le, rule ln_float_bounds(1)) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2505 | from mult_less_le_imp_less[OF \<open>real_of_float x < 2\<close> up] * low | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2506 | show "real_of_float ?max < 1" by (cases "real_of_float (lapprox_rat prec 2 3) = 0", | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2507 | auto simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2508 | show "0 \<le> real_of_float ?max" by (auto simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2509 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2510 | also have "\<dots> \<le> ln (real_of_float x * 2/3)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2511 | proof (rule ln_le_cancel_iff[symmetric, THEN iffD1]) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2512 | show "0 < real_of_float ?max + 1" by (auto simp add: real_of_float_max) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2513 | show "0 < real_of_float x * 2/3" using * by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2514 | show "real_of_float ?max + 1 \<le> real_of_float x * 2/3" using * up | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2515 | by (cases "0 < real_of_float x * real_of_float (lapprox_posrat prec 2 3) - 1", | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2516 | auto simp add: max_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2517 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2518 | finally have "?lb_horner (Float 1 (- 1)) + ?lb_horner ?max \<le> ln x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2519 | using ln_float_bounds(1)[of "Float 1 (- 1)" prec prec] add | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2520 | by (auto intro!: add_mono float_round_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2521 | note float_round_down_le[OF this, of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2522 | } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2523 | ultimately | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2524 | show ?thesis unfolding lb_ln.simps unfolding ub_ln.simps Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2525 | using \<open>\<not> x \<le> 0\<close> \<open>\<not> x < 1\<close> True False by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2526 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2527 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2528 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2529 | hence "\<not> x \<le> 0" and "\<not> x < 1" "0 < x" "\<not> x \<le> Float 3 (- 1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2530 | using \<open>1 \<le> x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2531 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2532 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2533 | define m where "m = mantissa x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2534 | define e where "e = exponent x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2535 | from Float_mantissa_exponent[of x] have Float: "x = Float m e" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2536 | by (simp add: m_def e_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2537 | let ?s = "Float (e + (bitlen m - 1)) 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2538 | let ?x = "Float m (- (bitlen m - 1))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2539 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2540 | have "0 < m" and "m \<noteq> 0" using \<open>0 < x\<close> Float powr_gt_zero[of 2 e] | 
| 67573 | 2541 | by (auto simp add: zero_less_mult_iff) | 
| 65582 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2542 | define bl where "bl = bitlen m - 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2543 | hence "bl \<ge> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2544 | using \<open>m > 0\<close> by (simp add: bitlen_alt_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2545 | have "1 \<le> Float m e" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2546 | using \<open>1 \<le> x\<close> Float unfolding less_eq_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2547 | from bitlen_div[OF \<open>0 < m\<close>] float_gt1_scale[OF \<open>1 \<le> Float m e\<close>] \<open>bl \<ge> 0\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2548 | have x_bnds: "0 \<le> real_of_float (?x - 1)" "real_of_float (?x - 1) < 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2549 | unfolding bl_def[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2550 | by (auto simp: powr_realpow[symmetric] field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2551 | (auto simp : powr_minus field_simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2552 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2553 |     {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2554 | have "float_round_down prec (lb_ln2 prec * ?s) \<le> ln 2 * (e + (bitlen m - 1))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2555 | (is "real_of_float ?lb2 \<le> _") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2556 | apply (rule float_round_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2557 | unfolding nat_0 power_0 mult_1_right times_float.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2558 | using lb_ln2[of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2559 | proof (rule mult_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2560 | from float_gt1_scale[OF \<open>1 \<le> Float m e\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2561 | show "0 \<le> real_of_float (Float (e + (bitlen m - 1)) 0)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2562 | qed auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2563 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2564 | from ln_float_bounds(1)[OF x_bnds] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2565 | have "float_round_down prec ((?x - 1) * lb_ln_horner prec (get_even prec) 1 (?x - 1)) \<le> ln ?x" (is "real_of_float ?lb_horner \<le> _") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2566 | by (auto intro!: float_round_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2567 | ultimately have "float_plus_down prec ?lb2 ?lb_horner \<le> ln x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2568 | unfolding Float ln_shifted_float[OF \<open>0 < m\<close>, of e] by (auto intro!: float_plus_down_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2569 | } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2570 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2571 |     {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2572 | from ln_float_bounds(2)[OF x_bnds] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2573 | have "ln ?x \<le> float_round_up prec ((?x - 1) * ub_ln_horner prec (get_odd prec) 1 (?x - 1))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2574 | (is "_ \<le> real_of_float ?ub_horner") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2575 | by (auto intro!: float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2576 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2577 | have "ln 2 * (e + (bitlen m - 1)) \<le> float_round_up prec (ub_ln2 prec * ?s)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2578 | (is "_ \<le> real_of_float ?ub2") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2579 | apply (rule float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2580 | unfolding nat_0 power_0 mult_1_right times_float.rep_eq | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2581 | using ub_ln2[of prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2582 | proof (rule mult_mono) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2583 | from float_gt1_scale[OF \<open>1 \<le> Float m e\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2584 | show "0 \<le> real_of_int (e + (bitlen m - 1))" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2585 | have "0 \<le> ln (2 :: real)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2586 | thus "0 \<le> real_of_float (ub_ln2 prec)" using ub_ln2[of prec] by arith | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2587 | qed auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2588 | ultimately have "ln x \<le> float_plus_up prec ?ub2 ?ub_horner" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2589 | unfolding Float ln_shifted_float[OF \<open>0 < m\<close>, of e] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2590 | by (auto intro!: float_plus_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2591 | } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2592 | ultimately show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2593 | unfolding lb_ln.simps | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2594 | unfolding ub_ln.simps | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2595 | unfolding if_not_P[OF \<open>\<not> x \<le> 0\<close>] if_not_P[OF \<open>\<not> x < 1\<close>] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2596 | if_not_P[OF False] if_not_P[OF \<open>\<not> x \<le> Float 3 (- 1)\<close>] Let_def | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2597 | unfolding plus_float.rep_eq e_def[symmetric] m_def[symmetric] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2598 | by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2599 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2600 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2601 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2602 | lemma ub_ln_lb_ln_bounds: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2603 | assumes "0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2604 | shows "the (lb_ln prec x) \<le> ln x \<and> ln x \<le> the (ub_ln prec x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2605 | (is "?lb \<le> ?ln \<and> ?ln \<le> ?ub") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2606 | proof (cases "x < 1") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2607 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2608 | hence "1 \<le> x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2609 | unfolding less_float_def less_eq_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2610 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2611 | using ub_ln_lb_ln_bounds'[OF \<open>1 \<le> x\<close>] . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2612 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2613 | case True | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2614 | have "\<not> x \<le> 0" using \<open>0 < x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2615 | from True have "real_of_float x \<le> 1" "x \<le> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2616 | by simp_all | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2617 | have "0 < real_of_float x" and "real_of_float x \<noteq> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2618 | using \<open>0 < x\<close> by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2619 | hence A: "0 < 1 / real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2620 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2621 |   {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2622 | let ?divl = "float_divl (max prec 1) 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2623 | have A': "1 \<le> ?divl" using float_divl_pos_less1_bound[OF \<open>0 < real_of_float x\<close> \<open>real_of_float x \<le> 1\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2624 | hence B: "0 < real_of_float ?divl" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2625 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2626 | have "ln ?divl \<le> ln (1 / x)" unfolding ln_le_cancel_iff[OF B A] using float_divl[of _ 1 x] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2627 | hence "ln x \<le> - ln ?divl" unfolding nonzero_inverse_eq_divide[OF \<open>real_of_float x \<noteq> 0\<close>, symmetric] ln_inverse[OF \<open>0 < real_of_float x\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2628 | from this ub_ln_lb_ln_bounds'[OF A', THEN conjunct1, THEN le_imp_neg_le] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2629 | have "?ln \<le> - the (lb_ln prec ?divl)" unfolding uminus_float.rep_eq by (rule order_trans) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2630 | } moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2631 |   {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2632 | let ?divr = "float_divr prec 1 x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2633 | have A': "1 \<le> ?divr" using float_divr_pos_less1_lower_bound[OF \<open>0 < x\<close> \<open>x \<le> 1\<close>] unfolding less_eq_float_def less_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2634 | hence B: "0 < real_of_float ?divr" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2635 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2636 | have "ln (1 / x) \<le> ln ?divr" unfolding ln_le_cancel_iff[OF A B] using float_divr[of 1 x] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2637 | hence "- ln ?divr \<le> ln x" unfolding nonzero_inverse_eq_divide[OF \<open>real_of_float x \<noteq> 0\<close>, symmetric] ln_inverse[OF \<open>0 < real_of_float x\<close>] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2638 | from ub_ln_lb_ln_bounds'[OF A', THEN conjunct2, THEN le_imp_neg_le] this | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2639 | have "- the (ub_ln prec ?divr) \<le> ?ln" unfolding uminus_float.rep_eq by (rule order_trans) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2640 | } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2641 | ultimately show ?thesis unfolding lb_ln.simps[where x=x] ub_ln.simps[where x=x] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2642 | unfolding if_not_P[OF \<open>\<not> x \<le> 0\<close>] if_P[OF True] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2643 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2644 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2645 | lemma lb_ln: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2646 | assumes "Some y = lb_ln prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2647 | shows "y \<le> ln x" and "0 < real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2648 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2649 | have "0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2650 | proof (rule ccontr) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2651 | assume "\<not> 0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2652 | hence "x \<le> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2653 | unfolding less_eq_float_def less_float_def by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2654 | thus False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2655 | using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2656 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2657 | thus "0 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2658 | have "the (lb_ln prec x) \<le> ln x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2659 | using ub_ln_lb_ln_bounds[OF \<open>0 < x\<close>] .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2660 | thus "y \<le> ln x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2661 | unfolding assms[symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2662 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2663 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2664 | lemma ub_ln: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2665 | assumes "Some y = ub_ln prec x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2666 | shows "ln x \<le> y" and "0 < real_of_float x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2667 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2668 | have "0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2669 | proof (rule ccontr) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2670 | assume "\<not> 0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2671 | hence "x \<le> 0" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2672 | thus False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2673 | using assms by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2674 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2675 | thus "0 < real_of_float x" by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2676 | have "ln x \<le> the (ub_ln prec x)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2677 | using ub_ln_lb_ln_bounds[OF \<open>0 < x\<close>] .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2678 | thus "ln x \<le> y" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2679 | unfolding assms[symmetric] by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2680 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2681 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2682 | lemma bnds_ln: "\<forall>(x::real) lx ux. (Some l, Some u) = | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2683 |   (lb_ln prec lx, ub_ln prec ux) \<and> x \<in> {lx .. ux} \<longrightarrow> l \<le> ln x \<and> ln x \<le> u"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2684 | proof (rule allI, rule allI, rule allI, rule impI) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2685 | fix x :: real | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2686 | fix lx ux | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2687 |   assume "(Some l, Some u) = (lb_ln prec lx, ub_ln prec ux) \<and> x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2688 |   hence l: "Some l = lb_ln prec lx " and u: "Some u = ub_ln prec ux" and x: "x \<in> {lx .. ux}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2689 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2690 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2691 | have "ln ux \<le> u" and "0 < real_of_float ux" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2692 | using ub_ln u by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2693 | have "l \<le> ln lx" and "0 < real_of_float lx" and "0 < x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2694 | using lb_ln[OF l] x by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2695 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2696 | from ln_le_cancel_iff[OF \<open>0 < real_of_float lx\<close> \<open>0 < x\<close>] \<open>l \<le> ln lx\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2697 | have "l \<le> ln x" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2698 | using x unfolding atLeastAtMost_iff by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2699 | moreover | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2700 | from ln_le_cancel_iff[OF \<open>0 < x\<close> \<open>0 < real_of_float ux\<close>] \<open>ln ux \<le> real_of_float u\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2701 | have "ln x \<le> u" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2702 | using x unfolding atLeastAtMost_iff by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2703 | ultimately show "l \<le> ln x \<and> ln x \<le> u" .. | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2704 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2705 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2706 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2707 | section \<open>Real power function\<close> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2708 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2709 | definition bnds_powr :: "nat \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float \<times> float) option" where | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2710 | "bnds_powr prec l1 u1 l2 u2 = ( | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2711 | if l1 = 0 \<and> u1 = 0 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2712 | Some (0, 0) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2713 | else if l1 = 0 \<and> l2 \<ge> 1 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2714 | let uln = the (ub_ln prec u1) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2715 | in Some (0, ub_exp prec (float_round_up prec (uln * (if uln \<ge> 0 then u2 else l2)))) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2716 | else if l1 \<le> 0 then | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2717 | None | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2718 | else | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2719 | Some (map_bnds lb_exp ub_exp prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2720 | (bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2721 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2722 | lemmas [simp del] = lb_exp.simps ub_exp.simps | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2723 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2724 | lemma mono_exp_real: "mono (exp :: real \<Rightarrow> real)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2725 | by (auto simp: mono_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2726 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2727 | lemma ub_exp_nonneg: "real_of_float (ub_exp prec x) \<ge> 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2728 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2729 | have "0 \<le> exp (real_of_float x)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2730 | also from exp_boundaries[of x prec] | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2731 | have "\<dots> \<le> real_of_float (ub_exp prec x)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2732 | finally show ?thesis . | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2733 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2734 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2735 | lemma bnds_powr: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2736 | assumes lu: "Some (l, u) = bnds_powr prec l1 u1 l2 u2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2737 |   assumes x: "x \<in> {real_of_float l1..real_of_float u1}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2738 |   assumes y: "y \<in> {real_of_float l2..real_of_float u2}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2739 |   shows   "x powr y \<in> {real_of_float l..real_of_float u}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2740 | proof - | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2741 | consider "l1 = 0" "u1 = 0" | "l1 = 0" "u1 \<noteq> 0" "l2 \<ge> 1" | | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2742 | "l1 \<le> 0" "\<not>(l1 = 0 \<and> (u1 = 0 \<or> l2 \<ge> 1))" | "l1 > 0" by force | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2743 | thus ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2744 | proof cases | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2745 | assume "l1 = 0" "u1 = 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2746 | with x lu show ?thesis by (auto simp: bnds_powr_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2747 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2748 | assume A: "l1 = 0" "u1 \<noteq> 0" "l2 \<ge> 1" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2749 | define uln where "uln = the (ub_ln prec u1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2750 | show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2751 | proof (cases "x = 0") | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2752 | case False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2753 | with A x y have "x powr y = exp (ln x * y)" by (simp add: powr_def) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2754 |       also {
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2755 | from A x False have "ln x \<le> ln (real_of_float u1)" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2756 | also from ub_ln_lb_ln_bounds[of u1 prec] A y x False | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2757 | have "ln (real_of_float u1) \<le> real_of_float uln" by (simp add: uln_def del: lb_ln.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2758 | also from A x y have "\<dots> * y \<le> real_of_float uln * (if uln \<ge> 0 then u2 else l2)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2759 | by (auto intro: mult_left_mono mult_left_mono_neg) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2760 | also have "\<dots> \<le> real_of_float (float_round_up prec (uln * (if uln \<ge> 0 then u2 else l2)))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2761 | by (simp add: float_round_up_le) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2762 | finally have "ln x * y \<le> \<dots>" using A y by - simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2763 | } | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2764 | also have "exp (real_of_float (float_round_up prec (uln * (if uln \<ge> 0 then u2 else l2)))) \<le> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2765 | real_of_float (ub_exp prec (float_round_up prec | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2766 | (uln * (if uln \<ge> 0 then u2 else l2))))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2767 | using exp_boundaries by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2768 | finally show ?thesis using A x y lu | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2769 | by (simp add: bnds_powr_def uln_def Let_def del: lb_ln.simps ub_ln.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2770 | qed (insert x y lu A, simp_all add: bnds_powr_def Let_def ub_exp_nonneg | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2771 | del: lb_ln.simps ub_ln.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2772 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2773 | assume "l1 \<le> 0" "\<not>(l1 = 0 \<and> (u1 = 0 \<or> l2 \<ge> 1))" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2774 | with lu show ?thesis by (simp add: bnds_powr_def split: if_split_asm) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2775 | next | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2776 | assume l1: "l1 > 0" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2777 | obtain lm um where lmum: | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2778 | "(lm, um) = bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2779 | by (cases "bnds_mult prec (the (lb_ln prec l1)) (the (ub_ln prec u1)) l2 u2") simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2780 | with l1 have "(l, u) = map_bnds lb_exp ub_exp prec (lm, um)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2781 | using lu by (simp add: bnds_powr_def del: lb_ln.simps ub_ln.simps split: if_split_asm) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2782 |     hence "exp (ln x * y) \<in> {real_of_float l..real_of_float u}"
 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2783 | proof (rule map_bnds[OF _ mono_exp_real], goal_cases) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2784 | case 1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2785 | let ?lln = "the (lb_ln prec l1)" and ?uln = "the (ub_ln prec u1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2786 | from ub_ln_lb_ln_bounds[of l1 prec] ub_ln_lb_ln_bounds[of u1 prec] x l1 | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2787 | have "real_of_float ?lln \<le> ln (real_of_float l1) \<and> | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2788 | ln (real_of_float u1) \<le> real_of_float ?uln" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2789 | by (auto simp del: lb_ln.simps ub_ln.simps) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2790 | moreover from l1 x have "ln (real_of_float l1) \<le> ln x \<and> ln x \<le> ln (real_of_float u1)" | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2791 | by auto | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2792 | ultimately have ln: "real_of_float ?lln \<le> ln x \<and> ln x \<le> real_of_float ?uln" by simp | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2793 | from lmum show ?case | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2794 | by (rule bnds_mult) (insert y ln, simp_all) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2795 | qed (insert exp_boundaries[of lm prec] exp_boundaries[of um prec], simp_all) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2796 | with x l1 show ?thesis | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2797 | by (simp add: powr_def mult_ac) | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2798 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2799 | qed | 
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2800 | |
| 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 eberlm <eberlm@in.tum.de> parents: diff
changeset | 2801 | end |