src/HOL/Decision_Procs/ex/Approximation_Ex.thy
author wenzelm
Sun, 30 Jan 2011 13:02:18 +0100
changeset 41648 6d736d983d5c
parent 32650 34bfa2492298
child 47600 e12289b5796b
permissions -rw-r--r--
clarified example settings for Proof General;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30443
873fa77be5f0 Extended approximation boundaries by fractions and base-2 floating point numbers
hoelzl
parents: 30429
diff changeset
     1
(* Author:     Johannes Hoelzl <hoelzl@in.tum.de> 2009 *)
30122
1c912a9d8200 standard headers;
wenzelm
parents: 29805
diff changeset
     2
30429
39acdf031548 moved Decision_Procs examples to Decision_Procs/ex
haftmann
parents: 30413
diff changeset
     3
theory Approximation_Ex
39acdf031548 moved Decision_Procs examples to Decision_Procs/ex
haftmann
parents: 30413
diff changeset
     4
imports Complex_Main "~~/src/HOL/Decision_Procs/Approximation"
29805
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
     5
begin
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
     6
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
     7
text {*
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
     8
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
     9
Here are some examples how to use the approximation method.
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    10
31863
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    11
The approximation method has the following syntax:
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    12
31864
90ec13cf7ab0 removed latex markup - there is no document generated from Decision_Procs/ex
hoelzl
parents: 31863
diff changeset
    13
approximate "prec" (splitting: "x" = "depth" and "y" = "depth" ...)? (taylor: "z" = "derivates")?
31863
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    14
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    15
Here "prec" specifies the precision used in all computations, it is specified as
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    16
number of bits to calculate. In the proposition to prove an arbitrary amount of
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    17
variables can be used, but each one need to be bounded by an upper and lower
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    18
bound.
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    19
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    20
To specify the bounds either @{term "l\<^isub>1 \<le> x \<and> x \<le> u\<^isub>1"},
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    21
@{term "x \<in> { l\<^isub>1 .. u\<^isub>1 }"} or @{term "x = bnd"} can be used. Where the
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    22
bound specification are again arithmetic formulas containing variables. They can
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    23
be connected using either meta level or HOL equivalence.
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    24
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    25
To use interval splitting add for each variable whos interval should be splitted
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    26
to the "splitting:" parameter. The parameter specifies how often each interval
31864
90ec13cf7ab0 removed latex markup - there is no document generated from Decision_Procs/ex
hoelzl
parents: 31863
diff changeset
    27
should be divided, e.g. when x = 16 is specified, there will be @{term "65536 = 2^16"}
90ec13cf7ab0 removed latex markup - there is no document generated from Decision_Procs/ex
hoelzl
parents: 31863
diff changeset
    28
intervals to be calculated.
31863
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    29
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    30
To use taylor series expansion specify the variable to derive. You need to
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    31
specify the amount of derivations to compute. When using taylor series expansion
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    32
only one variable can be used.
29805
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    33
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    34
*}
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    35
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    36
section "Compute some transcendental values"
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    37
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    38
lemma "\<bar> ln 2 - 544531980202654583340825686620847 / 785593587443817081832229725798400 \<bar> < inverse (2^51) "
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    39
  by (approximation 80)
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    40
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    41
lemma "\<bar> exp 1.626 - 5.083499996273 \<bar> < (inverse 10 ^ 10 :: real)"
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    42
  by (approximation 80)
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    43
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    44
lemma "\<bar> sqrt 2 - 1.4142135623730951 \<bar> < inverse 10 ^ 16"
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    45
  by (approximation 80)
31387
c4a3c3e9dc8e Removed usage of reference in reification
hoelzl
parents: 30443
diff changeset
    46
29805
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    47
lemma "\<bar> pi - 3.1415926535897932385 \<bar> < inverse 10 ^ 18"
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    48
  by (approximation 80)
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    49
31467
f7d2aa438bee Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
hoelzl
parents: 31387
diff changeset
    50
lemma "\<bar> sin 100 + 0.50636564110975879 \<bar> < inverse 10 ^ 17"
f7d2aa438bee Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
hoelzl
parents: 31387
diff changeset
    51
  by (approximation 80)
f7d2aa438bee Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
hoelzl
parents: 31387
diff changeset
    52
29805
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    53
section "Use variable ranges"
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    54
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    55
lemma "0.5 \<le> x \<and> x \<le> 4.5 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    56
  by (approximation 10)
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    57
31811
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    58
lemma "x \<in> {0.5 .. 4.5} \<longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    59
  by (approximation 10)
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    60
30413
c41afa5607be Fixed type error which appeared when Approximation bounds where specified as floating point numbers
hoelzl
parents: 30122
diff changeset
    61
lemma "0.49 \<le> x \<and> x \<le> 4.49 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
c41afa5607be Fixed type error which appeared when Approximation bounds where specified as floating point numbers
hoelzl
parents: 30122
diff changeset
    62
  by (approximation 20)
c41afa5607be Fixed type error which appeared when Approximation bounds where specified as floating point numbers
hoelzl
parents: 30122
diff changeset
    63
30443
873fa77be5f0 Extended approximation boundaries by fractions and base-2 floating point numbers
hoelzl
parents: 30429
diff changeset
    64
lemma "1 / 2^1 \<le> x \<and> x \<le> 9 / 2^1 \<Longrightarrow> \<bar> arctan x - 0.91 \<bar> < 0.455"
30413
c41afa5607be Fixed type error which appeared when Approximation bounds where specified as floating point numbers
hoelzl
parents: 30122
diff changeset
    65
  by (approximation 10)
c41afa5607be Fixed type error which appeared when Approximation bounds where specified as floating point numbers
hoelzl
parents: 30122
diff changeset
    66
31467
f7d2aa438bee Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
hoelzl
parents: 31387
diff changeset
    67
lemma "3.2 \<le> x \<and> x \<le> 6.2 \<Longrightarrow> sin x \<le> 0"
f7d2aa438bee Approximation: Implemented argument reduction for cosine. Sinus is now implemented in terms of cosine. Sqrt computes on the entire real numbers
hoelzl
parents: 31387
diff changeset
    68
  by (approximation 9)
29805
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    69
31811
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    70
lemma
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    71
  defines "g \<equiv> 9.80665" and "v \<equiv> 128.61" and "d \<equiv> pi / 180"
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    72
  shows "g / v * tan (35 * d) \<in> { 3 * d .. 3.1 * d }"
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    73
  using assms by (approximation 80)
64dea9a15031 Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents: 31467
diff changeset
    74
32650
34bfa2492298 correct variable order in approximate-method
hoelzl
parents: 31864
diff changeset
    75
lemma "x \<in> { 0 .. 1 :: real } \<longrightarrow> x ^ 2 \<le> x"
34bfa2492298 correct variable order in approximate-method
hoelzl
parents: 31864
diff changeset
    76
  by (approximation 30 splitting: x=1 taylor: x = 3)
34bfa2492298 correct variable order in approximate-method
hoelzl
parents: 31864
diff changeset
    77
34bfa2492298 correct variable order in approximate-method
hoelzl
parents: 31864
diff changeset
    78
value [approximate] "10"
31863
e391eee8bf14 Implemented taylor series expansion for approximation
hoelzl
parents: 31811
diff changeset
    79
29805
a5da150bd0ab Add approximation method
hoelzl
parents:
diff changeset
    80
end