John Harrison's example: a 32-bit approximation to pi. SLOW
authorpaulson <lp15@cam.ac.uk>
Wed, 01 Apr 2015 15:47:55 +0100
changeset 59871 e1a49ac9c537
parent 59870 68d6b6aa4450
child 59872 db4000b71fdb
John Harrison's example: a 32-bit approximation to pi. SLOW
src/HOL/ROOT
src/HOL/ex/Approximations.thy
src/HOL/ex/BinEx.thy
--- a/src/HOL/ROOT	Wed Apr 01 14:48:38 2015 +0100
+++ b/src/HOL/ROOT	Wed Apr 01 15:47:55 2015 +0100
@@ -527,6 +527,7 @@
     "~~/src/HOL/Library/Transitive_Closure_Table"
     Cartouche_Examples
   theories
+    Approximations
     Commands
     Adhoc_Overloading_Examples
     Iff_Oracle
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Approximations.thy	Wed Apr 01 15:47:55 2015 +0100
@@ -0,0 +1,39 @@
+section {* Binary Approximations to Constants *}
+
+theory Approximations
+imports "~~/src/HOL/Multivariate_Analysis/Complex_Transcendental"
+
+begin
+
+declare of_real_numeral [simp]
+
+subsection{*Approximation to pi*}
+
+lemma sin_pi6_straddle:
+  assumes "0 \<le> a" "a \<le> b" "b \<le> 4" "sin(a/6) \<le> 1/2" "1/2 \<le> sin(b/6)"
+    shows "a \<le> pi \<and> pi \<le> b"
+proof -
+  have *: "\<And>x::real. 0 < x & x < 7/5 \<Longrightarrow> 0 < sin x"
+    using pi_ge_two
+    by (auto intro: sin_gt_zero)
+  have ab: "(b \<le> pi * 3 \<Longrightarrow> pi \<le> b)" "(a \<le> pi * 3 \<Longrightarrow> a \<le> pi)"
+    using sin_mono_le_eq [of "pi/6" "b/6"] sin_mono_le_eq [of "a/6" "pi/6"] assms
+    by (simp_all add: sin_30 power.power_Suc norm_divide)
+  show ?thesis
+    using assms Taylor_sin [of "a/6" 0] pi_ge_two
+    by (auto simp: sin_30 power.power_Suc norm_divide intro: ab)
+qed
+
+(*32-bit approximation. SLOW simplification steps: big calculations with the rewriting engine*)
+lemma pi_approx_32: "abs(pi - 13493037705/4294967296) \<le> inverse(2 ^ 32)"
+  apply (simp only: abs_diff_le_iff)
+  apply (rule sin_pi6_straddle, simp_all)
+  using Taylor_sin [of "1686629713/3221225472" 11]
+  apply (simp add: in_Reals_norm sin_coeff_def Re_sin atMost_nat_numeral fact_numeral)
+  apply (simp only: pos_le_divide_eq [symmetric])
+  using Taylor_sin [of "6746518853/12884901888" 11]
+  apply (simp add: in_Reals_norm sin_coeff_def Re_sin atMost_nat_numeral fact_numeral)
+  apply (simp only: pos_le_divide_eq [symmetric] pos_divide_le_eq [symmetric])
+  done
+
+end
--- a/src/HOL/ex/BinEx.thy	Wed Apr 01 14:48:38 2015 +0100
+++ b/src/HOL/ex/BinEx.thy	Wed Apr 01 15:47:55 2015 +0100
@@ -78,6 +78,9 @@
 lemma "- (2*i) + 3  + (2*i + 4) = (0::int)"
 apply simp  oops
 
+(*Tobias's example dated 2015-03-02*)
+lemma "(pi * (real u * 2) = pi * (real (xa v) * - 2))"
+apply simp oops
 
 
 subsection {* Arithmetic Method Tests *}