| author | paulson <lp15@cam.ac.uk> | 
| Wed, 02 May 2018 12:47:56 +0100 | |
| changeset 68064 | b249fab48c76 | 
| parent 66453 | cc19f7ca2ed6 | 
| child 69529 | 4ab9657b3257 | 
| permissions | -rw-r--r-- | 
| 58834 | 1  | 
(* Author: Johannes Hoelzl, TU Muenchen  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
2  | 
Coercions removed by Dmitriy Traytel *)  | 
| 30122 | 3  | 
|
| 40892 | 4  | 
theory Approximation  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41024 
diff
changeset
 | 
5  | 
imports  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
41024 
diff
changeset
 | 
6  | 
Complex_Main  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
65582 
diff
changeset
 | 
7  | 
"HOL-Library.Code_Target_Numeral"  | 
| 
65582
 
a1bc1b020cf2
tuned Approximation: separated general material from oracle
 
eberlm <eberlm@in.tum.de> 
parents: 
65578 
diff
changeset
 | 
8  | 
Approximation_Bounds  | 
| 56923 | 9  | 
keywords "approximate" :: diag  | 
| 29805 | 10  | 
begin  | 
11  | 
||
12  | 
section "Implement floatarith"  | 
|
13  | 
||
14  | 
subsection "Define syntax and semantics"  | 
|
15  | 
||
| 58310 | 16  | 
datatype floatarith  | 
| 29805 | 17  | 
= Add floatarith floatarith  | 
18  | 
| Minus floatarith  | 
|
19  | 
| Mult floatarith floatarith  | 
|
20  | 
| Inverse floatarith  | 
|
21  | 
| Cos floatarith  | 
|
22  | 
| Arctan floatarith  | 
|
23  | 
| Abs floatarith  | 
|
24  | 
| Max floatarith floatarith  | 
|
25  | 
| Min floatarith floatarith  | 
|
26  | 
| Pi  | 
|
27  | 
| Sqrt floatarith  | 
|
28  | 
| Exp floatarith  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
29  | 
| Powr floatarith floatarith  | 
| 29805 | 30  | 
| Ln floatarith  | 
31  | 
| Power floatarith nat  | 
|
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
32  | 
| Floor floatarith  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
33  | 
| Var nat  | 
| 29805 | 34  | 
| Num float  | 
35  | 
||
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
36  | 
fun interpret_floatarith :: "floatarith \<Rightarrow> real list \<Rightarrow> real" where  | 
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
37  | 
"interpret_floatarith (Add a b) vs = (interpret_floatarith a vs) + (interpret_floatarith b vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
38  | 
"interpret_floatarith (Minus a) vs = - (interpret_floatarith a vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
39  | 
"interpret_floatarith (Mult a b) vs = (interpret_floatarith a vs) * (interpret_floatarith b vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
40  | 
"interpret_floatarith (Inverse a) vs = inverse (interpret_floatarith a vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
41  | 
"interpret_floatarith (Cos a) vs = cos (interpret_floatarith a vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
42  | 
"interpret_floatarith (Arctan a) vs = arctan (interpret_floatarith a vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
43  | 
"interpret_floatarith (Min a b) vs = min (interpret_floatarith a vs) (interpret_floatarith b vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
44  | 
"interpret_floatarith (Max a b) vs = max (interpret_floatarith a vs) (interpret_floatarith b vs)" |  | 
| 61945 | 45  | 
"interpret_floatarith (Abs a) vs = \<bar>interpret_floatarith a vs\<bar>" |  | 
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
46  | 
"interpret_floatarith Pi vs = pi" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
47  | 
"interpret_floatarith (Sqrt a) vs = sqrt (interpret_floatarith a vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
48  | 
"interpret_floatarith (Exp a) vs = exp (interpret_floatarith a vs)" |  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
49  | 
"interpret_floatarith (Powr a b) vs = interpret_floatarith a vs powr interpret_floatarith b vs" |  | 
| 
31098
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
50  | 
"interpret_floatarith (Ln a) vs = ln (interpret_floatarith a vs)" |  | 
| 
 
73dd67adf90a
replaced Ifloat => real_of_float and real, renamed ApproxEq => inequality, uneq => interpret_inequality, uneq' => approx_inequality, Ifloatarith => interpret_floatarith
 
hoelzl 
parents: 
30971 
diff
changeset
 | 
51  | 
"interpret_floatarith (Power a n) vs = (interpret_floatarith a vs)^n" |  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
52  | 
"interpret_floatarith (Floor a) vs = floor (interpret_floatarith a vs)" |  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
53  | 
"interpret_floatarith (Num f) vs = f" |  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
54  | 
"interpret_floatarith (Var n) vs = vs ! n"  | 
| 29805 | 55  | 
|
| 60680 | 56  | 
lemma interpret_floatarith_divide:  | 
57  | 
"interpret_floatarith (Mult a (Inverse b)) vs =  | 
|
58  | 
(interpret_floatarith a vs) / (interpret_floatarith b vs)"  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36531 
diff
changeset
 | 
59  | 
unfolding divide_inverse interpret_floatarith.simps ..  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
60  | 
|
| 60680 | 61  | 
lemma interpret_floatarith_diff:  | 
62  | 
"interpret_floatarith (Add a (Minus b)) vs =  | 
|
63  | 
(interpret_floatarith a vs) - (interpret_floatarith b vs)"  | 
|
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
53077 
diff
changeset
 | 
64  | 
unfolding interpret_floatarith.simps by simp  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
65  | 
|
| 60680 | 66  | 
lemma interpret_floatarith_sin:  | 
67  | 
"interpret_floatarith (Cos (Add (Mult Pi (Num (Float 1 (- 1)))) (Minus a))) vs =  | 
|
68  | 
sin (interpret_floatarith a vs)"  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
69  | 
unfolding sin_cos_eq interpret_floatarith.simps  | 
| 60680 | 70  | 
interpret_floatarith_divide interpret_floatarith_diff  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
71  | 
by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
72  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
73  | 
|
| 29805 | 74  | 
subsection "Implement approximation function"  | 
75  | 
||
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
76  | 
fun lift_bin :: "(float * float) option \<Rightarrow> (float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float * float) option) \<Rightarrow> (float * float) option" where  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
77  | 
"lift_bin (Some (l1, u1)) (Some (l2, u2)) f = f l1 u1 l2 u2" |  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
78  | 
"lift_bin a b f = None"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
79  | 
|
| 29805 | 80  | 
fun lift_bin' :: "(float * float) option \<Rightarrow> (float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> float \<Rightarrow> (float * float)) \<Rightarrow> (float * float) option" where  | 
81  | 
"lift_bin' (Some (l1, u1)) (Some (l2, u2)) f = Some (f l1 u1 l2 u2)" |  | 
|
82  | 
"lift_bin' a b f = None"  | 
|
83  | 
||
84  | 
fun lift_un :: "(float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> ((float option) * (float option))) \<Rightarrow> (float * float) option" where  | 
|
85  | 
"lift_un (Some (l1, u1)) f = (case (f l1 u1) of (Some l, Some u) \<Rightarrow> Some (l, u)  | 
|
86  | 
| t \<Rightarrow> None)" |  | 
|
87  | 
"lift_un b f = None"  | 
|
88  | 
||
89  | 
fun lift_un' :: "(float * float) option \<Rightarrow> (float \<Rightarrow> float \<Rightarrow> (float * float)) \<Rightarrow> (float * float) option" where  | 
|
90  | 
"lift_un' (Some (l1, u1)) f = Some (f l1 u1)" |  | 
|
91  | 
"lift_un' b f = None"  | 
|
92  | 
||
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
93  | 
definition bounded_by :: "real list \<Rightarrow> (float \<times> float) option list \<Rightarrow> bool" where  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
94  | 
"bounded_by xs vs \<longleftrightarrow>  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
95  | 
(\<forall> i < length vs. case vs ! i of None \<Rightarrow> True  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
96  | 
         | Some (l, u) \<Rightarrow> xs ! i \<in> { real_of_float l .. real_of_float u })"
 | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
97  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
98  | 
lemma bounded_byE:  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
99  | 
assumes "bounded_by xs vs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
100  | 
shows "\<And> i. i < length vs \<Longrightarrow> case vs ! i of None \<Rightarrow> True  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
101  | 
         | Some (l, u) \<Rightarrow> xs ! i \<in> { real_of_float l .. real_of_float u }"
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
102  | 
using assms bounded_by_def by blast  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
103  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
104  | 
lemma bounded_by_update:  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
105  | 
assumes "bounded_by xs vs"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
106  | 
    and bnd: "xs ! i \<in> { real_of_float l .. real_of_float u }"
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
107  | 
shows "bounded_by xs (vs[i := Some (l,u)])"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
108  | 
proof -  | 
| 60680 | 109  | 
  {
 | 
110  | 
fix j  | 
|
111  | 
let ?vs = "vs[i := Some (l,u)]"  | 
|
112  | 
assume "j < length ?vs"  | 
|
113  | 
hence [simp]: "j < length vs" by simp  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
114  | 
    have "case ?vs ! j of None \<Rightarrow> True | Some (l, u) \<Rightarrow> xs ! j \<in> { real_of_float l .. real_of_float u }"
 | 
| 60680 | 115  | 
proof (cases "?vs ! j")  | 
116  | 
case (Some b)  | 
|
117  | 
thus ?thesis  | 
|
118  | 
proof (cases "i = j")  | 
|
119  | 
case True  | 
|
120  | 
thus ?thesis using \<open>?vs ! j = Some b\<close> and bnd by auto  | 
|
121  | 
next  | 
|
122  | 
case False  | 
|
123  | 
thus ?thesis using \<open>bounded_by xs vs\<close> unfolding bounded_by_def by auto  | 
|
124  | 
qed  | 
|
125  | 
qed auto  | 
|
126  | 
}  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
127  | 
thus ?thesis unfolding bounded_by_def by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
128  | 
qed  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
129  | 
|
| 60680 | 130  | 
lemma bounded_by_None: "bounded_by xs (replicate (length xs) None)"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
131  | 
unfolding bounded_by_def by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
132  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
133  | 
fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option" where  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47108 
diff
changeset
 | 
134  | 
"approx' prec a bs = (case (approx prec a bs) of Some (l, u) \<Rightarrow> Some (float_round_down prec l, float_round_up prec u) | None \<Rightarrow> None)" |  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
135  | 
"approx prec (Add a b) bs =  | 
| 
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
136  | 
lift_bin' (approx' prec a bs) (approx' prec b bs)  | 
| 
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
137  | 
(\<lambda> l1 u1 l2 u2. (float_plus_down prec l1 l2, float_plus_up prec u1 u2))" |  | 
| 29805 | 138  | 
"approx prec (Minus a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (-u, -l))" |  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
139  | 
"approx prec (Mult a b) bs =  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
140  | 
lift_bin' (approx' prec a bs) (approx' prec b bs) (bnds_mult prec)" |  | 
| 29805 | 141  | 
"approx prec (Inverse a) bs = lift_un (approx' prec a bs) (\<lambda> l u. if (0 < l \<or> u < 0) then (Some (float_divl prec 1 u), Some (float_divr prec 1 l)) else (None, None))" |  | 
142  | 
"approx prec (Cos a) bs = lift_un' (approx' prec a bs) (bnds_cos prec)" |  | 
|
143  | 
"approx prec Pi bs = Some (lb_pi prec, ub_pi prec)" |  | 
|
144  | 
"approx prec (Min a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (min l1 l2, min u1 u2))" |  | 
|
145  | 
"approx prec (Max a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (\<lambda> l1 u1 l2 u2. (max l1 l2, max u1 u2))" |  | 
|
146  | 
"approx prec (Abs a) bs = lift_un' (approx' prec a bs) (\<lambda>l u. (if l < 0 \<and> 0 < u then 0 else min \<bar>l\<bar> \<bar>u\<bar>, max \<bar>l\<bar> \<bar>u\<bar>))" |  | 
|
147  | 
"approx prec (Arctan a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (lb_arctan prec l, ub_arctan prec u))" |  | 
|
| 
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: 
31148 
diff
changeset
 | 
148  | 
"approx prec (Sqrt a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (lb_sqrt prec l, ub_sqrt prec u))" |  | 
| 29805 | 149  | 
"approx prec (Exp a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (lb_exp prec l, ub_exp prec u))" |  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
150  | 
"approx prec (Powr a b) bs = lift_bin (approx' prec a bs) (approx' prec b bs) (bnds_powr prec)" |  | 
| 29805 | 151  | 
"approx prec (Ln a) bs = lift_un (approx' prec a bs) (\<lambda> l u. (lb_ln prec l, ub_ln prec u))" |  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
152  | 
"approx prec (Power a n) bs = lift_un' (approx' prec a bs) (float_power_bnds prec n)" |  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
153  | 
"approx prec (Floor a) bs = lift_un' (approx' prec a bs) (\<lambda> l u. (floor_fl l, floor_fl u))" |  | 
| 29805 | 154  | 
"approx prec (Num f) bs = Some (f, f)" |  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
155  | 
"approx prec (Var i) bs = (if i < length bs then bs ! i else None)"  | 
| 29805 | 156  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
157  | 
lemma approx_approx':  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
158  | 
assumes Pa: "\<And>l u. Some (l, u) = approx prec a vs \<Longrightarrow>  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
159  | 
l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
160  | 
and approx': "Some (l, u) = approx' prec a vs"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
161  | 
shows "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
162  | 
proof -  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
163  | 
obtain l' u' where S: "Some (l', u') = approx prec a vs"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
164  | 
using approx' unfolding approx'.simps by (cases "approx prec a vs") auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
165  | 
have l': "l = float_round_down prec l'" and u': "u = float_round_up prec u'"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
166  | 
using approx' unfolding approx'.simps S[symmetric] by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
167  | 
show ?thesis unfolding l' u'  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
168  | 
using order_trans[OF Pa[OF S, THEN conjunct2] float_round_up[of u']]  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
169  | 
using order_trans[OF float_round_down[of _ l'] Pa[OF S, THEN conjunct1]] by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
170  | 
qed  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
171  | 
|
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
172  | 
lemma lift_bin_ex:  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
173  | 
assumes lift_bin_Some: "Some (l, u) = lift_bin a b f"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
174  | 
shows "\<exists> l1 u1 l2 u2. Some (l1, u1) = a \<and> Some (l2, u2) = b"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
175  | 
proof (cases a)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
176  | 
case None  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
177  | 
hence "None = lift_bin a b f"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
178  | 
unfolding None lift_bin.simps ..  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
179  | 
thus ?thesis  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
180  | 
using lift_bin_Some by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
181  | 
next  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
182  | 
case (Some a')  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
183  | 
show ?thesis  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
184  | 
proof (cases b)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
185  | 
case None  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
186  | 
hence "None = lift_bin a b f"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
187  | 
unfolding None lift_bin.simps ..  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
188  | 
thus ?thesis using lift_bin_Some by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
189  | 
next  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
190  | 
case (Some b')  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
191  | 
obtain la ua where a': "a' = (la, ua)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
192  | 
by (cases a') auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
193  | 
obtain lb ub where b': "b' = (lb, ub)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
194  | 
by (cases b') auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
195  | 
thus ?thesis  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
196  | 
unfolding \<open>a = Some a'\<close> \<open>b = Some b'\<close> a' b' by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
197  | 
qed  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
198  | 
qed  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
199  | 
|
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
200  | 
lemma lift_bin_f:  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
201  | 
assumes lift_bin_Some: "Some (l, u) = lift_bin (g a) (g b) f"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
202  | 
and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
203  | 
and Pb: "\<And>l u. Some (l, u) = g b \<Longrightarrow> P l u b"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
204  | 
shows "\<exists> l1 u1 l2 u2. P l1 u1 a \<and> P l2 u2 b \<and> Some (l, u) = f l1 u1 l2 u2"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
205  | 
proof -  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
206  | 
obtain l1 u1 l2 u2  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
207  | 
where Sa: "Some (l1, u1) = g a"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
208  | 
and Sb: "Some (l2, u2) = g b"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
209  | 
using lift_bin_ex[OF assms(1)] by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
210  | 
have lu: "Some (l, u) = f l1 u1 l2 u2"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
211  | 
using lift_bin_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin.simps] by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
212  | 
thus ?thesis  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
213  | 
using Pa[OF Sa] Pb[OF Sb] by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
214  | 
qed  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
215  | 
|
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
216  | 
lemma lift_bin:  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
217  | 
assumes lift_bin_Some: "Some (l, u) = lift_bin (approx' prec a bs) (approx' prec b bs) f"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
218  | 
and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
219  | 
real_of_float l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real_of_float u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
220  | 
and Pb: "\<And>l u. Some (l, u) = approx prec b bs \<Longrightarrow>  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
221  | 
real_of_float l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real_of_float u"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
222  | 
shows "\<exists>l1 u1 l2 u2. (real_of_float l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> real_of_float u1) \<and>  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
223  | 
(real_of_float l2 \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> real_of_float u2) \<and>  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
224  | 
Some (l, u) = (f l1 u1 l2 u2)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
225  | 
proof -  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
226  | 
  { fix l u assume "Some (l, u) = approx' prec a bs"
 | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
227  | 
with approx_approx'[of prec a bs, OF _ this] Pa  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
228  | 
have "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u" by auto } note Pa = this  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
229  | 
  { fix l u assume "Some (l, u) = approx' prec b bs"
 | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
230  | 
with approx_approx'[of prec b bs, OF _ this] Pb  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
231  | 
have "l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u" by auto } note Pb = this  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
232  | 
|
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
233  | 
from lift_bin_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_bin_Some, OF Pa Pb]  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
234  | 
show ?thesis by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
235  | 
qed  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
236  | 
|
| 29805 | 237  | 
lemma lift_bin'_ex:  | 
238  | 
assumes lift_bin'_Some: "Some (l, u) = lift_bin' a b f"  | 
|
239  | 
shows "\<exists> l1 u1 l2 u2. Some (l1, u1) = a \<and> Some (l2, u2) = b"  | 
|
240  | 
proof (cases a)  | 
|
| 60680 | 241  | 
case None  | 
242  | 
hence "None = lift_bin' a b f"  | 
|
243  | 
unfolding None lift_bin'.simps ..  | 
|
244  | 
thus ?thesis  | 
|
245  | 
using lift_bin'_Some by auto  | 
|
| 29805 | 246  | 
next  | 
247  | 
case (Some a')  | 
|
248  | 
show ?thesis  | 
|
249  | 
proof (cases b)  | 
|
| 60680 | 250  | 
case None  | 
251  | 
hence "None = lift_bin' a b f"  | 
|
252  | 
unfolding None lift_bin'.simps ..  | 
|
| 29805 | 253  | 
thus ?thesis using lift_bin'_Some by auto  | 
254  | 
next  | 
|
255  | 
case (Some b')  | 
|
| 60680 | 256  | 
obtain la ua where a': "a' = (la, ua)"  | 
257  | 
by (cases a') auto  | 
|
258  | 
obtain lb ub where b': "b' = (lb, ub)"  | 
|
259  | 
by (cases b') auto  | 
|
260  | 
thus ?thesis  | 
|
261  | 
unfolding \<open>a = Some a'\<close> \<open>b = Some b'\<close> a' b' by auto  | 
|
| 29805 | 262  | 
qed  | 
263  | 
qed  | 
|
264  | 
||
265  | 
lemma lift_bin'_f:  | 
|
266  | 
assumes lift_bin'_Some: "Some (l, u) = lift_bin' (g a) (g b) f"  | 
|
| 60680 | 267  | 
and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"  | 
268  | 
and Pb: "\<And>l u. Some (l, u) = g b \<Longrightarrow> P l u b"  | 
|
| 29805 | 269  | 
shows "\<exists> l1 u1 l2 u2. P l1 u1 a \<and> P l2 u2 b \<and> l = fst (f l1 u1 l2 u2) \<and> u = snd (f l1 u1 l2 u2)"  | 
270  | 
proof -  | 
|
271  | 
obtain l1 u1 l2 u2  | 
|
| 60680 | 272  | 
where Sa: "Some (l1, u1) = g a"  | 
273  | 
and Sb: "Some (l2, u2) = g b"  | 
|
274  | 
using lift_bin'_ex[OF assms(1)] by auto  | 
|
275  | 
have lu: "(l, u) = f l1 u1 l2 u2"  | 
|
276  | 
using lift_bin'_Some[unfolded Sa[symmetric] Sb[symmetric] lift_bin'.simps] by auto  | 
|
277  | 
have "l = fst (f l1 u1 l2 u2)" and "u = snd (f l1 u1 l2 u2)"  | 
|
278  | 
unfolding lu[symmetric] by auto  | 
|
279  | 
thus ?thesis  | 
|
280  | 
using Pa[OF Sa] Pb[OF Sb] by auto  | 
|
| 29805 | 281  | 
qed  | 
282  | 
||
283  | 
lemma lift_bin':  | 
|
284  | 
assumes lift_bin'_Some: "Some (l, u) = lift_bin' (approx' prec a bs) (approx' prec b bs) f"  | 
|
| 60680 | 285  | 
and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>  | 
286  | 
l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u" (is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")  | 
|
287  | 
and Pb: "\<And>l u. Some (l, u) = approx prec b bs \<Longrightarrow>  | 
|
288  | 
l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u"  | 
|
289  | 
shows "\<exists>l1 u1 l2 u2. (l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u1) \<and>  | 
|
290  | 
(l2 \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u2) \<and>  | 
|
291  | 
l = fst (f l1 u1 l2 u2) \<and> u = snd (f l1 u1 l2 u2)"  | 
|
| 29805 | 292  | 
proof -  | 
293  | 
  { fix l u assume "Some (l, u) = approx' prec a bs"
 | 
|
294  | 
with approx_approx'[of prec a bs, OF _ this] Pa  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
295  | 
have "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u" by auto } note Pa = this  | 
| 29805 | 296  | 
  { fix l u assume "Some (l, u) = approx' prec b bs"
 | 
297  | 
with approx_approx'[of prec b bs, OF _ this] Pb  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
298  | 
have "l \<le> interpret_floatarith b xs \<and> interpret_floatarith b xs \<le> u" by auto } note Pb = this  | 
| 29805 | 299  | 
|
300  | 
from lift_bin'_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_bin'_Some, OF Pa Pb]  | 
|
301  | 
show ?thesis by auto  | 
|
302  | 
qed  | 
|
303  | 
||
304  | 
lemma lift_un'_ex:  | 
|
305  | 
assumes lift_un'_Some: "Some (l, u) = lift_un' a f"  | 
|
306  | 
shows "\<exists> l u. Some (l, u) = a"  | 
|
307  | 
proof (cases a)  | 
|
| 60680 | 308  | 
case None  | 
309  | 
hence "None = lift_un' a f"  | 
|
310  | 
unfolding None lift_un'.simps ..  | 
|
311  | 
thus ?thesis  | 
|
312  | 
using lift_un'_Some by auto  | 
|
| 29805 | 313  | 
next  | 
314  | 
case (Some a')  | 
|
| 60680 | 315  | 
obtain la ua where a': "a' = (la, ua)"  | 
316  | 
by (cases a') auto  | 
|
317  | 
thus ?thesis  | 
|
318  | 
unfolding \<open>a = Some a'\<close> a' by auto  | 
|
| 29805 | 319  | 
qed  | 
320  | 
||
321  | 
lemma lift_un'_f:  | 
|
322  | 
assumes lift_un'_Some: "Some (l, u) = lift_un' (g a) f"  | 
|
| 60680 | 323  | 
and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"  | 
| 29805 | 324  | 
shows "\<exists> l1 u1. P l1 u1 a \<and> l = fst (f l1 u1) \<and> u = snd (f l1 u1)"  | 
325  | 
proof -  | 
|
| 60680 | 326  | 
obtain l1 u1 where Sa: "Some (l1, u1) = g a"  | 
327  | 
using lift_un'_ex[OF assms(1)] by auto  | 
|
328  | 
have lu: "(l, u) = f l1 u1"  | 
|
329  | 
using lift_un'_Some[unfolded Sa[symmetric] lift_un'.simps] by auto  | 
|
330  | 
have "l = fst (f l1 u1)" and "u = snd (f l1 u1)"  | 
|
331  | 
unfolding lu[symmetric] by auto  | 
|
332  | 
thus ?thesis  | 
|
333  | 
using Pa[OF Sa] by auto  | 
|
| 29805 | 334  | 
qed  | 
335  | 
||
336  | 
lemma lift_un':  | 
|
337  | 
assumes lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f"  | 
|
| 60680 | 338  | 
and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>  | 
339  | 
l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
|
340  | 
(is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")  | 
|
341  | 
shows "\<exists>l1 u1. (l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u1) \<and>  | 
|
342  | 
l = fst (f l1 u1) \<and> u = snd (f l1 u1)"  | 
|
| 29805 | 343  | 
proof -  | 
| 60680 | 344  | 
have Pa: "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
345  | 
if "Some (l, u) = approx' prec a bs" for l u  | 
|
346  | 
using approx_approx'[of prec a bs, OF _ that] Pa  | 
|
347  | 
by auto  | 
|
| 29805 | 348  | 
from lift_un'_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_un'_Some, OF Pa]  | 
349  | 
show ?thesis by auto  | 
|
350  | 
qed  | 
|
351  | 
||
352  | 
lemma lift_un'_bnds:  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
353  | 
  assumes bnds: "\<forall> (x::real) lx ux. (l, u) = f lx ux \<and> x \<in> { lx .. ux } \<longrightarrow> l \<le> f' x \<and> f' x \<le> u"
 | 
| 60680 | 354  | 
and lift_un'_Some: "Some (l, u) = lift_un' (approx' prec a bs) f"  | 
355  | 
and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>  | 
|
356  | 
l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
357  | 
shows "real_of_float l \<le> f' (interpret_floatarith a xs) \<and> f' (interpret_floatarith a xs) \<le> real_of_float u"  | 
| 29805 | 358  | 
proof -  | 
359  | 
from lift_un'[OF lift_un'_Some Pa]  | 
|
| 60680 | 360  | 
obtain l1 u1 where "l1 \<le> interpret_floatarith a xs"  | 
361  | 
and "interpret_floatarith a xs \<le> u1"  | 
|
362  | 
and "l = fst (f l1 u1)"  | 
|
363  | 
and "u = snd (f l1 u1)"  | 
|
364  | 
by blast  | 
|
365  | 
  hence "(l, u) = f l1 u1" and "interpret_floatarith a xs \<in> {l1 .. u1}"
 | 
|
366  | 
by auto  | 
|
367  | 
thus ?thesis  | 
|
368  | 
using bnds by auto  | 
|
| 29805 | 369  | 
qed  | 
370  | 
||
371  | 
lemma lift_un_ex:  | 
|
372  | 
assumes lift_un_Some: "Some (l, u) = lift_un a f"  | 
|
| 60680 | 373  | 
shows "\<exists>l u. Some (l, u) = a"  | 
| 29805 | 374  | 
proof (cases a)  | 
| 60680 | 375  | 
case None  | 
376  | 
hence "None = lift_un a f"  | 
|
377  | 
unfolding None lift_un.simps ..  | 
|
378  | 
thus ?thesis  | 
|
379  | 
using lift_un_Some by auto  | 
|
| 29805 | 380  | 
next  | 
381  | 
case (Some a')  | 
|
| 60680 | 382  | 
obtain la ua where a': "a' = (la, ua)"  | 
383  | 
by (cases a') auto  | 
|
384  | 
thus ?thesis  | 
|
385  | 
unfolding \<open>a = Some a'\<close> a' by auto  | 
|
| 29805 | 386  | 
qed  | 
387  | 
||
388  | 
lemma lift_un_f:  | 
|
389  | 
assumes lift_un_Some: "Some (l, u) = lift_un (g a) f"  | 
|
| 60680 | 390  | 
and Pa: "\<And>l u. Some (l, u) = g a \<Longrightarrow> P l u a"  | 
| 29805 | 391  | 
shows "\<exists> l1 u1. P l1 u1 a \<and> Some l = fst (f l1 u1) \<and> Some u = snd (f l1 u1)"  | 
392  | 
proof -  | 
|
| 60680 | 393  | 
obtain l1 u1 where Sa: "Some (l1, u1) = g a"  | 
394  | 
using lift_un_ex[OF assms(1)] by auto  | 
|
| 29805 | 395  | 
have "fst (f l1 u1) \<noteq> None \<and> snd (f l1 u1) \<noteq> None"  | 
396  | 
proof (rule ccontr)  | 
|
397  | 
assume "\<not> (fst (f l1 u1) \<noteq> None \<and> snd (f l1 u1) \<noteq> None)"  | 
|
398  | 
hence or: "fst (f l1 u1) = None \<or> snd (f l1 u1) = None" by auto  | 
|
| 31809 | 399  | 
hence "lift_un (g a) f = None"  | 
| 29805 | 400  | 
proof (cases "fst (f l1 u1) = None")  | 
401  | 
case True  | 
|
| 60680 | 402  | 
then obtain b where b: "f l1 u1 = (None, b)"  | 
403  | 
by (cases "f l1 u1") auto  | 
|
404  | 
thus ?thesis  | 
|
405  | 
unfolding Sa[symmetric] lift_un.simps b by auto  | 
|
| 29805 | 406  | 
next  | 
| 60680 | 407  | 
case False  | 
408  | 
hence "snd (f l1 u1) = None"  | 
|
409  | 
using or by auto  | 
|
410  | 
with False obtain b where b: "f l1 u1 = (Some b, None)"  | 
|
411  | 
by (cases "f l1 u1") auto  | 
|
412  | 
thus ?thesis  | 
|
413  | 
unfolding Sa[symmetric] lift_un.simps b by auto  | 
|
| 29805 | 414  | 
qed  | 
| 60680 | 415  | 
thus False  | 
416  | 
using lift_un_Some by auto  | 
|
| 29805 | 417  | 
qed  | 
| 60680 | 418  | 
then obtain a' b' where f: "f l1 u1 = (Some a', Some b')"  | 
419  | 
by (cases "f l1 u1") auto  | 
|
| 29805 | 420  | 
from lift_un_Some[unfolded Sa[symmetric] lift_un.simps f]  | 
| 60680 | 421  | 
have "Some l = fst (f l1 u1)" and "Some u = snd (f l1 u1)"  | 
422  | 
unfolding f by auto  | 
|
423  | 
thus ?thesis  | 
|
424  | 
unfolding Sa[symmetric] lift_un.simps using Pa[OF Sa] by auto  | 
|
| 29805 | 425  | 
qed  | 
426  | 
||
427  | 
lemma lift_un:  | 
|
428  | 
assumes lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f"  | 
|
| 60680 | 429  | 
and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>  | 
430  | 
l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
|
431  | 
(is "\<And>l u. _ = ?g a \<Longrightarrow> ?P l u a")  | 
|
432  | 
shows "\<exists>l1 u1. (l1 \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u1) \<and>  | 
|
| 29805 | 433  | 
Some l = fst (f l1 u1) \<and> Some u = snd (f l1 u1)"  | 
434  | 
proof -  | 
|
| 60680 | 435  | 
have Pa: "l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
436  | 
if "Some (l, u) = approx' prec a bs" for l u  | 
|
437  | 
using approx_approx'[of prec a bs, OF _ that] Pa by auto  | 
|
| 29805 | 438  | 
from lift_un_f[where g="\<lambda>a. approx' prec a bs" and P = ?P, OF lift_un_Some, OF Pa]  | 
439  | 
show ?thesis by auto  | 
|
440  | 
qed  | 
|
441  | 
||
442  | 
lemma lift_un_bnds:  | 
|
| 60680 | 443  | 
  assumes bnds: "\<forall>(x::real) lx ux. (Some l, Some u) = f lx ux \<and> x \<in> { lx .. ux } \<longrightarrow> l \<le> f' x \<and> f' x \<le> u"
 | 
444  | 
and lift_un_Some: "Some (l, u) = lift_un (approx' prec a bs) f"  | 
|
445  | 
and Pa: "\<And>l u. Some (l, u) = approx prec a bs \<Longrightarrow>  | 
|
446  | 
l \<le> interpret_floatarith a xs \<and> interpret_floatarith a xs \<le> u"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
447  | 
shows "real_of_float l \<le> f' (interpret_floatarith a xs) \<and> f' (interpret_floatarith a xs) \<le> real_of_float u"  | 
| 29805 | 448  | 
proof -  | 
449  | 
from lift_un[OF lift_un_Some Pa]  | 
|
| 60680 | 450  | 
obtain l1 u1 where "l1 \<le> interpret_floatarith a xs"  | 
451  | 
and "interpret_floatarith a xs \<le> u1"  | 
|
452  | 
and "Some l = fst (f l1 u1)"  | 
|
453  | 
and "Some u = snd (f l1 u1)"  | 
|
454  | 
by blast  | 
|
455  | 
  hence "(Some l, Some u) = f l1 u1" and "interpret_floatarith a xs \<in> {l1 .. u1}"
 | 
|
456  | 
by auto  | 
|
457  | 
thus ?thesis  | 
|
458  | 
using bnds by auto  | 
|
| 29805 | 459  | 
qed  | 
460  | 
||
461  | 
lemma approx:  | 
|
462  | 
assumes "bounded_by xs vs"  | 
|
| 60680 | 463  | 
and "Some (l, u) = approx prec arith vs" (is "_ = ?g arith")  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
464  | 
shows "l \<le> interpret_floatarith arith xs \<and> interpret_floatarith arith xs \<le> u" (is "?P l u arith")  | 
| 60533 | 465  | 
using \<open>Some (l, u) = approx prec arith vs\<close>  | 
| 
45129
 
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
 
wenzelm 
parents: 
44821 
diff
changeset
 | 
466  | 
proof (induct arith arbitrary: l u)  | 
| 29805 | 467  | 
case (Add a b)  | 
468  | 
from lift_bin'[OF Add.prems[unfolded approx.simps]] Add.hyps  | 
|
| 60680 | 469  | 
obtain l1 u1 l2 u2 where "l = float_plus_down prec l1 l2"  | 
470  | 
and "u = float_plus_up prec u1 u2" "l1 \<le> interpret_floatarith a xs"  | 
|
471  | 
and "interpret_floatarith a xs \<le> u1" "l2 \<le> interpret_floatarith b xs"  | 
|
472  | 
and "interpret_floatarith b xs \<le> u2"  | 
|
473  | 
unfolding fst_conv snd_conv by blast  | 
|
474  | 
thus ?case  | 
|
475  | 
unfolding interpret_floatarith.simps by (auto intro!: float_plus_up_le float_plus_down_le)  | 
|
| 29805 | 476  | 
next  | 
477  | 
case (Minus a)  | 
|
478  | 
from lift_un'[OF Minus.prems[unfolded approx.simps]] Minus.hyps  | 
|
| 60680 | 479  | 
obtain l1 u1 where "l = -u1" "u = -l1"  | 
480  | 
and "l1 \<le> interpret_floatarith a xs" "interpret_floatarith a xs \<le> u1"  | 
|
481  | 
unfolding fst_conv snd_conv by blast  | 
|
482  | 
thus ?case  | 
|
483  | 
unfolding interpret_floatarith.simps using minus_float.rep_eq by auto  | 
|
| 29805 | 484  | 
next  | 
485  | 
case (Mult a b)  | 
|
486  | 
from lift_bin'[OF Mult.prems[unfolded approx.simps]] Mult.hyps  | 
|
| 31809 | 487  | 
obtain l1 u1 l2 u2  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
488  | 
where l: "l = fst (bnds_mult prec l1 u1 l2 u2)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
489  | 
and u: "u = snd (bnds_mult prec l1 u1 l2 u2)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
490  | 
and a: "l1 \<le> interpret_floatarith a xs" "interpret_floatarith a xs \<le> u1"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
491  | 
and b: "l2 \<le> interpret_floatarith b xs" "interpret_floatarith b xs \<le> u2" unfolding fst_conv snd_conv by blast  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
492  | 
from l u have lu: "(l, u) = bnds_mult prec l1 u1 l2 u2" by simp  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
493  | 
from bnds_mult[OF lu] a b show ?case by simp  | 
| 29805 | 494  | 
next  | 
495  | 
case (Inverse a)  | 
|
496  | 
from lift_un[OF Inverse.prems[unfolded approx.simps], unfolded if_distrib[of fst] if_distrib[of snd] fst_conv snd_conv] Inverse.hyps  | 
|
| 31809 | 497  | 
obtain l1 u1 where l': "Some l = (if 0 < l1 \<or> u1 < 0 then Some (float_divl prec 1 u1) else None)"  | 
| 29805 | 498  | 
and u': "Some u = (if 0 < l1 \<or> u1 < 0 then Some (float_divr prec 1 l1) else None)"  | 
| 60680 | 499  | 
and l1: "l1 \<le> interpret_floatarith a xs"  | 
500  | 
and u1: "interpret_floatarith a xs \<le> u1"  | 
|
501  | 
by blast  | 
|
502  | 
have either: "0 < l1 \<or> u1 < 0"  | 
|
503  | 
proof (rule ccontr)  | 
|
504  | 
assume P: "\<not> (0 < l1 \<or> u1 < 0)"  | 
|
505  | 
show False  | 
|
506  | 
using l' unfolding if_not_P[OF P] by auto  | 
|
507  | 
qed  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
508  | 
moreover have l1_le_u1: "real_of_float l1 \<le> real_of_float u1"  | 
| 60680 | 509  | 
using l1 u1 by auto  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
510  | 
ultimately have "real_of_float l1 \<noteq> 0" and "real_of_float u1 \<noteq> 0"  | 
| 60680 | 511  | 
by auto  | 
| 29805 | 512  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
513  | 
have inv: "inverse u1 \<le> inverse (interpret_floatarith a xs)  | 
| 
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
514  | 
\<and> inverse (interpret_floatarith a xs) \<le> inverse l1"  | 
| 29805 | 515  | 
proof (cases "0 < l1")  | 
| 60680 | 516  | 
case True  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
517  | 
hence "0 < real_of_float u1" and "0 < real_of_float l1" "0 < interpret_floatarith a xs"  | 
| 47600 | 518  | 
using l1_le_u1 l1 by auto  | 
| 29805 | 519  | 
show ?thesis  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
520  | 
unfolding inverse_le_iff_le[OF \<open>0 < real_of_float u1\<close> \<open>0 < interpret_floatarith a xs\<close>]  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
521  | 
inverse_le_iff_le[OF \<open>0 < interpret_floatarith a xs\<close> \<open>0 < real_of_float l1\<close>]  | 
| 29805 | 522  | 
using l1 u1 by auto  | 
523  | 
next  | 
|
| 60680 | 524  | 
case False  | 
525  | 
hence "u1 < 0"  | 
|
526  | 
using either by blast  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
527  | 
hence "real_of_float u1 < 0" and "real_of_float l1 < 0" "interpret_floatarith a xs < 0"  | 
| 47600 | 528  | 
using l1_le_u1 u1 by auto  | 
| 29805 | 529  | 
show ?thesis  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
530  | 
unfolding inverse_le_iff_le_neg[OF \<open>real_of_float u1 < 0\<close> \<open>interpret_floatarith a xs < 0\<close>]  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
531  | 
inverse_le_iff_le_neg[OF \<open>interpret_floatarith a xs < 0\<close> \<open>real_of_float l1 < 0\<close>]  | 
| 29805 | 532  | 
using l1 u1 by auto  | 
533  | 
qed  | 
|
| 
31468
 
b8267feaf342
Approximation: Corrected precision of ln on all real values
 
hoelzl 
parents: 
31467 
diff
changeset
 | 
534  | 
|
| 60680 | 535  | 
from l' have "l = float_divl prec 1 u1"  | 
536  | 
by (cases "0 < l1 \<or> u1 < 0") auto  | 
|
537  | 
hence "l \<le> inverse u1"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
538  | 
unfolding nonzero_inverse_eq_divide[OF \<open>real_of_float u1 \<noteq> 0\<close>]  | 
| 60680 | 539  | 
using float_divl[of prec 1 u1] by auto  | 
540  | 
also have "\<dots> \<le> inverse (interpret_floatarith a xs)"  | 
|
541  | 
using inv by auto  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
542  | 
finally have "l \<le> inverse (interpret_floatarith a xs)" .  | 
| 29805 | 543  | 
moreover  | 
| 60680 | 544  | 
from u' have "u = float_divr prec 1 l1"  | 
545  | 
by (cases "0 < l1 \<or> u1 < 0") auto  | 
|
546  | 
hence "inverse l1 \<le> u"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
547  | 
unfolding nonzero_inverse_eq_divide[OF \<open>real_of_float l1 \<noteq> 0\<close>]  | 
| 60680 | 548  | 
using float_divr[of 1 l1 prec] by auto  | 
549  | 
hence "inverse (interpret_floatarith a xs) \<le> u"  | 
|
550  | 
by (rule order_trans[OF inv[THEN conjunct2]])  | 
|
551  | 
ultimately show ?case  | 
|
552  | 
unfolding interpret_floatarith.simps using l1 u1 by auto  | 
|
| 29805 | 553  | 
next  | 
554  | 
case (Abs x)  | 
|
555  | 
from lift_un'[OF Abs.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Abs.hyps  | 
|
| 60680 | 556  | 
obtain l1 u1 where l': "l = (if l1 < 0 \<and> 0 < u1 then 0 else min \<bar>l1\<bar> \<bar>u1\<bar>)"  | 
557  | 
and u': "u = max \<bar>l1\<bar> \<bar>u1\<bar>"  | 
|
558  | 
and l1: "l1 \<le> interpret_floatarith x xs"  | 
|
559  | 
and u1: "interpret_floatarith x xs \<le> u1"  | 
|
560  | 
by blast  | 
|
561  | 
thus ?case  | 
|
562  | 
unfolding l' u'  | 
|
563  | 
by (cases "l1 < 0 \<and> 0 < u1") (auto simp add: real_of_float_min real_of_float_max)  | 
|
| 29805 | 564  | 
next  | 
565  | 
case (Min a b)  | 
|
566  | 
from lift_bin'[OF Min.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Min.hyps  | 
|
567  | 
obtain l1 u1 l2 u2 where l': "l = min l1 l2" and u': "u = min u1 u2"  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
568  | 
and l1: "l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> u1"  | 
| 60680 | 569  | 
and l1: "l2 \<le> interpret_floatarith b xs" and u1: "interpret_floatarith b xs \<le> u2"  | 
570  | 
by blast  | 
|
571  | 
thus ?case  | 
|
572  | 
unfolding l' u' by (auto simp add: real_of_float_min)  | 
|
| 29805 | 573  | 
next  | 
574  | 
case (Max a b)  | 
|
575  | 
from lift_bin'[OF Max.prems[unfolded approx.simps], unfolded fst_conv snd_conv] Max.hyps  | 
|
576  | 
obtain l1 u1 l2 u2 where l': "l = max l1 l2" and u': "u = max u1 u2"  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
577  | 
and l1: "l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> u1"  | 
| 60680 | 578  | 
and l1: "l2 \<le> interpret_floatarith b xs" and u1: "interpret_floatarith b xs \<le> u2"  | 
579  | 
by blast  | 
|
580  | 
thus ?case  | 
|
581  | 
unfolding l' u' by (auto simp add: real_of_float_max)  | 
|
582  | 
next  | 
|
583  | 
case (Cos a)  | 
|
584  | 
with lift_un'_bnds[OF bnds_cos] show ?case by auto  | 
|
585  | 
next  | 
|
586  | 
case (Arctan a)  | 
|
587  | 
with lift_un'_bnds[OF bnds_arctan] show ?case by auto  | 
|
588  | 
next  | 
|
589  | 
case Pi  | 
|
590  | 
with pi_boundaries show ?case by auto  | 
|
591  | 
next  | 
|
592  | 
case (Sqrt a)  | 
|
593  | 
with lift_un'_bnds[OF bnds_sqrt] show ?case by auto  | 
|
594  | 
next  | 
|
595  | 
case (Exp a)  | 
|
596  | 
with lift_un'_bnds[OF bnds_exp] show ?case by auto  | 
|
597  | 
next  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
598  | 
case (Powr a b)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
599  | 
from lift_bin[OF Powr.prems[unfolded approx.simps]] Powr.hyps  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
600  | 
obtain l1 u1 l2 u2 where lu: "Some (l, u) = bnds_powr prec l1 u1 l2 u2"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
601  | 
and l1: "l1 \<le> interpret_floatarith a xs" and u1: "interpret_floatarith a xs \<le> u1"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
602  | 
and l2: "l2 \<le> interpret_floatarith b xs" and u2: "interpret_floatarith b xs \<le> u2"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
603  | 
by blast  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
604  | 
from bnds_powr[OF lu] l1 u1 l2 u2  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
605  | 
show ?case by simp  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
606  | 
next  | 
| 60680 | 607  | 
case (Ln a)  | 
608  | 
with lift_un_bnds[OF bnds_ln] show ?case by auto  | 
|
609  | 
next  | 
|
610  | 
case (Power a n)  | 
|
611  | 
with lift_un'_bnds[OF bnds_power] show ?case by auto  | 
|
612  | 
next  | 
|
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
613  | 
case (Floor a)  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
614  | 
from lift_un'[OF Floor.prems[unfolded approx.simps] Floor.hyps]  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
615  | 
show ?case by (auto simp: floor_fl.rep_eq floor_mono)  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
616  | 
next  | 
| 60680 | 617  | 
case (Num f)  | 
618  | 
thus ?case by auto  | 
|
| 29805 | 619  | 
next  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
620  | 
case (Var n)  | 
| 60533 | 621  | 
from this[symmetric] \<open>bounded_by xs vs\<close>[THEN bounded_byE, of n]  | 
| 60680 | 622  | 
show ?case by (cases "n < length vs") auto  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
623  | 
qed  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
624  | 
|
| 58310 | 625  | 
datatype form = Bound floatarith floatarith floatarith form  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
626  | 
| Assign floatarith floatarith form  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
627  | 
| Less floatarith floatarith  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
628  | 
| LessEqual floatarith floatarith  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
629  | 
| AtLeastAtMost floatarith floatarith floatarith  | 
| 58986 | 630  | 
| Conj form form  | 
631  | 
| Disj form form  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
632  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
633  | 
fun interpret_form :: "form \<Rightarrow> real list \<Rightarrow> bool" where  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
634  | 
"interpret_form (Bound x a b f) vs = (interpret_floatarith x vs \<in> { interpret_floatarith a vs .. interpret_floatarith b vs } \<longrightarrow> interpret_form f vs)" |
 | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
635  | 
"interpret_form (Assign x a f) vs = (interpret_floatarith x vs = interpret_floatarith a vs \<longrightarrow> interpret_form f vs)" |  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
636  | 
"interpret_form (Less a b) vs = (interpret_floatarith a vs < interpret_floatarith b vs)" |  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
637  | 
"interpret_form (LessEqual a b) vs = (interpret_floatarith a vs \<le> interpret_floatarith b vs)" |  | 
| 58986 | 638  | 
"interpret_form (AtLeastAtMost x a b) vs = (interpret_floatarith x vs \<in> { interpret_floatarith a vs .. interpret_floatarith b vs })" |
 | 
639  | 
"interpret_form (Conj f g) vs \<longleftrightarrow> interpret_form f vs \<and> interpret_form g vs" |  | 
|
640  | 
"interpret_form (Disj f g) vs \<longleftrightarrow> interpret_form f vs \<or> interpret_form g vs"  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
641  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
642  | 
fun approx_form' and approx_form :: "nat \<Rightarrow> form \<Rightarrow> (float * float) option list \<Rightarrow> nat list \<Rightarrow> bool" where  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
643  | 
"approx_form' prec f 0 n l u bs ss = approx_form prec f (bs[n := Some (l, u)]) ss" |  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
644  | 
"approx_form' prec f (Suc s) n l u bs ss =  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
645  | 
(let m = (l + u) * Float 1 (- 1)  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
646  | 
in (if approx_form' prec f s n l m bs ss then approx_form' prec f s n m u bs ss else False))" |  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
647  | 
"approx_form prec (Bound (Var n) a b f) bs ss =  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
648  | 
(case (approx prec a bs, approx prec b bs)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
649  | 
of (Some (l, _), Some (_, u)) \<Rightarrow> approx_form' prec f (ss ! n) n l u bs ss  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
650  | 
| _ \<Rightarrow> False)" |  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
651  | 
"approx_form prec (Assign (Var n) a f) bs ss =  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
652  | 
(case (approx prec a bs)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
653  | 
of (Some (l, u)) \<Rightarrow> approx_form' prec f (ss ! n) n l u bs ss  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
654  | 
| _ \<Rightarrow> False)" |  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
655  | 
"approx_form prec (Less a b) bs ss =  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
656  | 
(case (approx prec a bs, approx prec b bs)  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
657  | 
of (Some (l, u), Some (l', u')) \<Rightarrow> float_plus_up prec u (-l') < 0  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
658  | 
| _ \<Rightarrow> False)" |  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
659  | 
"approx_form prec (LessEqual a b) bs ss =  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
660  | 
(case (approx prec a bs, approx prec b bs)  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
661  | 
of (Some (l, u), Some (l', u')) \<Rightarrow> float_plus_up prec u (-l') \<le> 0  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
662  | 
| _ \<Rightarrow> False)" |  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
663  | 
"approx_form prec (AtLeastAtMost x a b) bs ss =  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
664  | 
(case (approx prec x bs, approx prec a bs, approx prec b bs)  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
665  | 
of (Some (lx, ux), Some (l, u), Some (l', u')) \<Rightarrow> float_plus_up prec u (-lx) \<le> 0 \<and> float_plus_up prec ux (-l') \<le> 0  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
666  | 
| _ \<Rightarrow> False)" |  | 
| 58986 | 667  | 
"approx_form prec (Conj a b) bs ss \<longleftrightarrow> approx_form prec a bs ss \<and> approx_form prec b bs ss" |  | 
668  | 
"approx_form prec (Disj a b) bs ss \<longleftrightarrow> approx_form prec a bs ss \<or> approx_form prec b bs ss" |  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
669  | 
"approx_form _ _ _ _ = False"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
670  | 
|
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
671  | 
lemma lazy_conj: "(if A then B else False) = (A \<and> B)" by simp  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
672  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
673  | 
lemma approx_form_approx_form':  | 
| 60680 | 674  | 
assumes "approx_form' prec f s n l u bs ss"  | 
675  | 
    and "(x::real) \<in> { l .. u }"
 | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
676  | 
  obtains l' u' where "x \<in> { l' .. u' }"
 | 
| 49351 | 677  | 
and "approx_form prec f (bs[n := Some (l', u')]) ss"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
678  | 
using assms proof (induct s arbitrary: l u)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
679  | 
case 0  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
680  | 
from this(1)[of l u] this(2,3)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
681  | 
show thesis by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
682  | 
next  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
683  | 
case (Suc s)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
684  | 
|
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
685  | 
let ?m = "(l + u) * Float 1 (- 1)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
686  | 
have "real_of_float l \<le> ?m" and "?m \<le> real_of_float u"  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47108 
diff
changeset
 | 
687  | 
unfolding less_eq_float_def using Suc.prems by auto  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
688  | 
|
| 60533 | 689  | 
  with \<open>x \<in> { l .. u }\<close>
 | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
690  | 
  have "x \<in> { l .. ?m} \<or> x \<in> { ?m .. u }" by auto
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
691  | 
thus thesis  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
692  | 
proof (rule disjE)  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
693  | 
    assume *: "x \<in> { l .. ?m }"
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
694  | 
with Suc.hyps[OF _ _ *] Suc.prems  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
695  | 
show thesis by (simp add: Let_def lazy_conj)  | 
| 29805 | 696  | 
next  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
697  | 
    assume *: "x \<in> { ?m .. u }"
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
698  | 
with Suc.hyps[OF _ _ *] Suc.prems  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
699  | 
show thesis by (simp add: Let_def lazy_conj)  | 
| 29805 | 700  | 
qed  | 
701  | 
qed  | 
|
702  | 
||
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
703  | 
lemma approx_form_aux:  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
704  | 
assumes "approx_form prec f vs ss"  | 
| 49351 | 705  | 
and "bounded_by xs vs"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
706  | 
shows "interpret_form f xs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
707  | 
using assms proof (induct f arbitrary: vs)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
708  | 
case (Bound x a b f)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
709  | 
then obtain n  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
710  | 
where x_eq: "x = Var n" by (cases x) auto  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
711  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
712  | 
with Bound.prems obtain l u' l' u  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
713  | 
where l_eq: "Some (l, u') = approx prec a vs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
714  | 
and u_eq: "Some (l', u) = approx prec b vs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
715  | 
and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"  | 
| 
37411
 
c88c44156083
removed simplifier congruence rule of "prod_case"
 
haftmann 
parents: 
37391 
diff
changeset
 | 
716  | 
by (cases "approx prec a vs", simp) (cases "approx prec b vs", auto)  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
717  | 
|
| 60680 | 718  | 
have "interpret_form f xs"  | 
719  | 
    if "xs ! n \<in> { interpret_floatarith a xs .. interpret_floatarith b xs }"
 | 
|
720  | 
proof -  | 
|
721  | 
from approx[OF Bound.prems(2) l_eq] and approx[OF Bound.prems(2) u_eq] that  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
722  | 
    have "xs ! n \<in> { l .. u}" by auto
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
723  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
724  | 
from approx_form_approx_form'[OF approx_form' this]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
725  | 
    obtain lx ux where bnds: "xs ! n \<in> { lx .. ux }"
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
726  | 
and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
727  | 
|
| 60680 | 728  | 
from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some (lx, ux)])"  | 
729  | 
by (rule bounded_by_update)  | 
|
730  | 
with Bound.hyps[OF approx_form] show ?thesis  | 
|
731  | 
by blast  | 
|
732  | 
qed  | 
|
733  | 
thus ?case  | 
|
734  | 
using interpret_form.simps x_eq and interpret_floatarith.simps by simp  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
735  | 
next  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
736  | 
case (Assign x a f)  | 
| 60680 | 737  | 
then obtain n where x_eq: "x = Var n"  | 
738  | 
by (cases x) auto  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
739  | 
|
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47108 
diff
changeset
 | 
740  | 
with Assign.prems obtain l u  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
741  | 
where bnd_eq: "Some (l, u) = approx prec a vs"  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
742  | 
and x_eq: "x = Var n"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
743  | 
and approx_form': "approx_form' prec f (ss ! n) n l u vs ss"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
744  | 
by (cases "approx prec a vs") auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
745  | 
|
| 60680 | 746  | 
have "interpret_form f xs"  | 
747  | 
if bnds: "xs ! n = interpret_floatarith a xs"  | 
|
748  | 
proof -  | 
|
749  | 
from approx[OF Assign.prems(2) bnd_eq] bnds  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
750  | 
    have "xs ! n \<in> { l .. u}" by auto
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
751  | 
from approx_form_approx_form'[OF approx_form' this]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
752  | 
    obtain lx ux where bnds: "xs ! n \<in> { lx .. ux }"
 | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
753  | 
and approx_form: "approx_form prec f (vs[n := Some (lx, ux)]) ss" .  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
754  | 
|
| 60680 | 755  | 
from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some (lx, ux)])"  | 
756  | 
by (rule bounded_by_update)  | 
|
757  | 
with Assign.hyps[OF approx_form] show ?thesis  | 
|
758  | 
by blast  | 
|
759  | 
qed  | 
|
760  | 
thus ?case  | 
|
761  | 
using interpret_form.simps x_eq and interpret_floatarith.simps by simp  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
762  | 
next  | 
| 29805 | 763  | 
case (Less a b)  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
764  | 
then obtain l u l' u'  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
765  | 
where l_eq: "Some (l, u) = approx prec a vs"  | 
| 49351 | 766  | 
and u_eq: "Some (l', u') = approx prec b vs"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
767  | 
and inequality: "real_of_float (float_plus_up prec u (-l')) < 0"  | 
| 60680 | 768  | 
by (cases "approx prec a vs", auto, cases "approx prec b vs", auto)  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
769  | 
from le_less_trans[OF float_plus_up inequality]  | 
| 
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
770  | 
approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq]  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
771  | 
show ?case by auto  | 
| 29805 | 772  | 
next  | 
773  | 
case (LessEqual a b)  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
774  | 
then obtain l u l' u'  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
775  | 
where l_eq: "Some (l, u) = approx prec a vs"  | 
| 49351 | 776  | 
and u_eq: "Some (l', u') = approx prec b vs"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
777  | 
and inequality: "real_of_float (float_plus_up prec u (-l')) \<le> 0"  | 
| 60680 | 778  | 
by (cases "approx prec a vs", auto, cases "approx prec b vs", auto)  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
779  | 
from order_trans[OF float_plus_up inequality]  | 
| 
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
780  | 
approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq]  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
781  | 
show ?case by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
782  | 
next  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
783  | 
case (AtLeastAtMost x a b)  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
784  | 
then obtain lx ux l u l' u'  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
785  | 
where x_eq: "Some (lx, ux) = approx prec x vs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
786  | 
and l_eq: "Some (l, u) = approx prec a vs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
787  | 
and u_eq: "Some (l', u') = approx prec b vs"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
788  | 
and inequality: "real_of_float (float_plus_up prec u (-lx)) \<le> 0" "real_of_float (float_plus_up prec ux (-l')) \<le> 0"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
789  | 
by (cases "approx prec x vs", auto,  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
790  | 
cases "approx prec a vs", auto,  | 
| 
56073
 
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
 
nipkow 
parents: 
55506 
diff
changeset
 | 
791  | 
cases "approx prec b vs", auto)  | 
| 
58985
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
792  | 
from order_trans[OF float_plus_up inequality(1)] order_trans[OF float_plus_up inequality(2)]  | 
| 
 
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
 
immler 
parents: 
58982 
diff
changeset
 | 
793  | 
approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq]  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
794  | 
show ?case by auto  | 
| 58986 | 795  | 
qed auto  | 
| 29805 | 796  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
797  | 
lemma approx_form:  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
798  | 
assumes "n = length xs"  | 
| 60680 | 799  | 
and "approx_form prec f (replicate n None) ss"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
800  | 
shows "interpret_form f xs"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
801  | 
using approx_form_aux[OF _ bounded_by_None] assms by auto  | 
| 29805 | 802  | 
|
| 60680 | 803  | 
|
| 60533 | 804  | 
subsection \<open>Implementing Taylor series expansion\<close>  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
805  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
806  | 
fun isDERIV :: "nat \<Rightarrow> floatarith \<Rightarrow> real list \<Rightarrow> bool" where  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
807  | 
"isDERIV x (Add a b) vs = (isDERIV x a vs \<and> isDERIV x b vs)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
808  | 
"isDERIV x (Mult a b) vs = (isDERIV x a vs \<and> isDERIV x b vs)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
809  | 
"isDERIV x (Minus a) vs = isDERIV x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
810  | 
"isDERIV x (Inverse a) vs = (isDERIV x a vs \<and> interpret_floatarith a vs \<noteq> 0)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
811  | 
"isDERIV x (Cos a) vs = isDERIV x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
812  | 
"isDERIV x (Arctan a) vs = isDERIV x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
813  | 
"isDERIV x (Min a b) vs = False" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
814  | 
"isDERIV x (Max a b) vs = False" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
815  | 
"isDERIV x (Abs a) vs = False" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
816  | 
"isDERIV x Pi vs = True" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
817  | 
"isDERIV x (Sqrt a) vs = (isDERIV x a vs \<and> interpret_floatarith a vs > 0)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
818  | 
"isDERIV x (Exp a) vs = isDERIV x a vs" |  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
819  | 
"isDERIV x (Powr a b) vs =  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
820  | 
(isDERIV x a vs \<and> isDERIV x b vs \<and> interpret_floatarith a vs > 0)" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
821  | 
"isDERIV x (Ln a) vs = (isDERIV x a vs \<and> interpret_floatarith a vs > 0)" |  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
822  | 
"isDERIV x (Floor a) vs = (isDERIV x a vs \<and> interpret_floatarith a vs \<notin> \<int>)" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
823  | 
"isDERIV x (Power a 0) vs = True" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
824  | 
"isDERIV x (Power a (Suc n)) vs = isDERIV x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
825  | 
"isDERIV x (Num f) vs = True" |  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
826  | 
"isDERIV x (Var n) vs = True"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
827  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
828  | 
fun DERIV_floatarith :: "nat \<Rightarrow> floatarith \<Rightarrow> floatarith" where  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
829  | 
"DERIV_floatarith x (Add a b) = Add (DERIV_floatarith x a) (DERIV_floatarith x b)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
830  | 
"DERIV_floatarith x (Mult a b) = Add (Mult a (DERIV_floatarith x b)) (Mult (DERIV_floatarith x a) b)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
831  | 
"DERIV_floatarith x (Minus a) = Minus (DERIV_floatarith x a)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
832  | 
"DERIV_floatarith x (Inverse a) = Minus (Mult (DERIV_floatarith x a) (Inverse (Power a 2)))" |  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
833  | 
"DERIV_floatarith x (Cos a) = Minus (Mult (Cos (Add (Mult Pi (Num (Float 1 (- 1)))) (Minus a))) (DERIV_floatarith x a))" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
834  | 
"DERIV_floatarith x (Arctan a) = Mult (Inverse (Add (Num 1) (Power a 2))) (DERIV_floatarith x a)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
835  | 
"DERIV_floatarith x (Min a b) = Num 0" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
836  | 
"DERIV_floatarith x (Max a b) = Num 0" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
837  | 
"DERIV_floatarith x (Abs a) = Num 0" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
838  | 
"DERIV_floatarith x Pi = Num 0" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
839  | 
"DERIV_floatarith x (Sqrt a) = (Mult (Inverse (Mult (Sqrt a) (Num 2))) (DERIV_floatarith x a))" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
840  | 
"DERIV_floatarith x (Exp a) = Mult (Exp a) (DERIV_floatarith x a)" |  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
841  | 
"DERIV_floatarith x (Powr a b) =  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
842  | 
Mult (Powr a b) (Add (Mult (DERIV_floatarith x b) (Ln a)) (Mult (Mult (DERIV_floatarith x a) b) (Inverse a)))" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
843  | 
"DERIV_floatarith x (Ln a) = Mult (Inverse a) (DERIV_floatarith x a)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
844  | 
"DERIV_floatarith x (Power a 0) = Num 0" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
845  | 
"DERIV_floatarith x (Power a (Suc n)) = Mult (Num (Float (int (Suc n)) 0)) (Mult (Power a n) (DERIV_floatarith x a))" |  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
846  | 
"DERIV_floatarith x (Floor a) = Num 0" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
847  | 
"DERIV_floatarith x (Num f) = Num 0" |  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
848  | 
"DERIV_floatarith x (Var n) = (if x = n then Num 1 else Num 0)"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
849  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
850  | 
lemma has_real_derivative_powr':  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
851  | 
fixes f g :: "real \<Rightarrow> real"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
852  | 
assumes "(f has_real_derivative f') (at x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
853  | 
assumes "(g has_real_derivative g') (at x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
854  | 
assumes "f x > 0"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
855  | 
defines "h \<equiv> \<lambda>x. f x powr g x * (g' * ln (f x) + f' * g x / f x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
856  | 
shows "((\<lambda>x. f x powr g x) has_real_derivative h x) (at x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
857  | 
proof (subst DERIV_cong_ev[OF refl _ refl])  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
858  | 
from assms have "isCont f x"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
859  | 
by (simp add: DERIV_continuous)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
860  | 
hence "f \<midarrow>x\<rightarrow> f x" by (simp add: continuous_at)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
861  | 
with \<open>f x > 0\<close> have "eventually (\<lambda>x. f x > 0) (nhds x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
862  | 
by (auto simp: tendsto_at_iff_tendsto_nhds dest: order_tendstoD)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
863  | 
thus "eventually (\<lambda>x. f x powr g x = exp (g x * ln (f x))) (nhds x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
864  | 
by eventually_elim (simp add: powr_def)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
865  | 
next  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
866  | 
from assms show "((\<lambda>x. exp (g x * ln (f x))) has_real_derivative h x) (at x)"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
867  | 
by (auto intro!: derivative_eq_intros simp: h_def powr_def)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
868  | 
qed  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
869  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
870  | 
lemma DERIV_floatarith:  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
871  | 
assumes "n < length vs"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
872  | 
assumes isDERIV: "isDERIV n f (vs[n := x])"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
873  | 
shows "DERIV (\<lambda> x'. interpret_floatarith f (vs[n := x'])) x :>  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
874  | 
interpret_floatarith (DERIV_floatarith n f) (vs[n := x])"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
875  | 
(is "DERIV (?i f) x :> _")  | 
| 49351 | 876  | 
using isDERIV  | 
877  | 
proof (induct f arbitrary: x)  | 
|
878  | 
case (Inverse a)  | 
|
879  | 
thus ?case  | 
|
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56195 
diff
changeset
 | 
880  | 
by (auto intro!: derivative_eq_intros simp add: algebra_simps power2_eq_square)  | 
| 49351 | 881  | 
next  | 
882  | 
case (Cos a)  | 
|
883  | 
thus ?case  | 
|
| 56382 | 884  | 
by (auto intro!: derivative_eq_intros  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
885  | 
simp del: interpret_floatarith.simps(5)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
886  | 
simp add: interpret_floatarith_sin interpret_floatarith.simps(5)[of a])  | 
| 49351 | 887  | 
next  | 
888  | 
case (Power a n)  | 
|
889  | 
thus ?case  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
890  | 
by (cases n) (auto intro!: derivative_eq_intros simp del: power_Suc)  | 
| 49351 | 891  | 
next  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
892  | 
case (Floor a)  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
893  | 
thus ?case  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
894  | 
by (auto intro!: derivative_eq_intros DERIV_isCont floor_has_real_derivative)  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
895  | 
next  | 
| 49351 | 896  | 
case (Ln a)  | 
| 56382 | 897  | 
thus ?case by (auto intro!: derivative_eq_intros simp add: divide_inverse)  | 
| 49351 | 898  | 
next  | 
899  | 
case (Var i)  | 
|
| 60533 | 900  | 
thus ?case using \<open>n < length vs\<close> by auto  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
901  | 
next  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
902  | 
case (Powr a b)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
903  | 
note [derivative_intros] = has_real_derivative_powr'  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
904  | 
from Powr show ?case  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
905  | 
by (auto intro!: derivative_eq_intros simp: field_simps)  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56195 
diff
changeset
 | 
906  | 
qed (auto intro!: derivative_eq_intros)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
907  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
908  | 
declare approx.simps[simp del]  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
909  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
910  | 
fun isDERIV_approx :: "nat \<Rightarrow> nat \<Rightarrow> floatarith \<Rightarrow> (float * float) option list \<Rightarrow> bool" where  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
911  | 
"isDERIV_approx prec x (Add a b) vs = (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
912  | 
"isDERIV_approx prec x (Mult a b) vs = (isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
913  | 
"isDERIV_approx prec x (Minus a) vs = isDERIV_approx prec x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
914  | 
"isDERIV_approx prec x (Inverse a) vs =  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
915  | 
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l \<or> u < 0 | None \<Rightarrow> False))" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
916  | 
"isDERIV_approx prec x (Cos a) vs = isDERIV_approx prec x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
917  | 
"isDERIV_approx prec x (Arctan a) vs = isDERIV_approx prec x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
918  | 
"isDERIV_approx prec x (Min a b) vs = False" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
919  | 
"isDERIV_approx prec x (Max a b) vs = False" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
920  | 
"isDERIV_approx prec x (Abs a) vs = False" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
921  | 
"isDERIV_approx prec x Pi vs = True" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
922  | 
"isDERIV_approx prec x (Sqrt a) vs =  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
923  | 
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
924  | 
"isDERIV_approx prec x (Exp a) vs = isDERIV_approx prec x a vs" |  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
925  | 
"isDERIV_approx prec x (Powr a b) vs =  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
926  | 
(isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
927  | 
"isDERIV_approx prec x (Ln a) vs =  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
928  | 
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> 0 < l | None \<Rightarrow> False))" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
929  | 
"isDERIV_approx prec x (Power a 0) vs = True" |  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
930  | 
"isDERIV_approx prec x (Floor a) vs =  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
931  | 
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some (l, u) \<Rightarrow> l > floor u \<and> u < ceiling l | None \<Rightarrow> False))" |  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
932  | 
"isDERIV_approx prec x (Power a (Suc n)) vs = isDERIV_approx prec x a vs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
933  | 
"isDERIV_approx prec x (Num f) vs = True" |  | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
934  | 
"isDERIV_approx prec x (Var n) vs = True"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
935  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
936  | 
lemma isDERIV_approx:  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
937  | 
assumes "bounded_by xs vs"  | 
| 49351 | 938  | 
and isDERIV_approx: "isDERIV_approx prec x f vs"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
939  | 
shows "isDERIV x f xs"  | 
| 49351 | 940  | 
using isDERIV_approx  | 
941  | 
proof (induct f)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
942  | 
case (Inverse a)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
943  | 
then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
944  | 
and *: "0 < l \<or> u < 0"  | 
| 49351 | 945  | 
by (cases "approx prec a vs") auto  | 
| 60533 | 946  | 
with approx[OF \<open>bounded_by xs vs\<close> approx_Some]  | 
| 47600 | 947  | 
have "interpret_floatarith a xs \<noteq> 0" by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
948  | 
thus ?case using Inverse by auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
949  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
950  | 
case (Ln a)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
951  | 
then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
952  | 
and *: "0 < l"  | 
| 49351 | 953  | 
by (cases "approx prec a vs") auto  | 
| 60533 | 954  | 
with approx[OF \<open>bounded_by xs vs\<close> approx_Some]  | 
| 47600 | 955  | 
have "0 < interpret_floatarith a xs" by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
956  | 
thus ?case using Ln by auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
957  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
958  | 
case (Sqrt a)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
959  | 
then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
960  | 
and *: "0 < l"  | 
| 49351 | 961  | 
by (cases "approx prec a vs") auto  | 
| 60533 | 962  | 
with approx[OF \<open>bounded_by xs vs\<close> approx_Some]  | 
| 47600 | 963  | 
have "0 < interpret_floatarith a xs" by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
964  | 
thus ?case using Sqrt by auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
965  | 
next  | 
| 60680 | 966  | 
case (Power a n)  | 
967  | 
thus ?case by (cases n) auto  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
968  | 
next  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
969  | 
case (Powr a b)  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
970  | 
from Powr obtain l1 u1 where a: "Some (l1, u1) = approx prec a vs" and pos: "0 < l1"  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
971  | 
by (cases "approx prec a vs") auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
972  | 
with approx[OF \<open>bounded_by xs vs\<close> a]  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
973  | 
have "0 < interpret_floatarith a xs" by auto  | 
| 
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
974  | 
with Powr show ?case by auto  | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
975  | 
next  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
976  | 
case (Floor a)  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
977  | 
then obtain l u where approx_Some: "Some (l, u) = approx prec a vs"  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
978  | 
and "real_of_int \<lfloor>real_of_float u\<rfloor> < real_of_float l" "real_of_float u < real_of_int \<lceil>real_of_float l\<rceil>"  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
979  | 
and "isDERIV x a xs"  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
980  | 
by (cases "approx prec a vs") auto  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
981  | 
with approx[OF \<open>bounded_by xs vs\<close> approx_Some] le_floor_iff  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
982  | 
show ?case  | 
| 
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
983  | 
by (force elim!: Ints_cases)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
984  | 
qed auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
985  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
986  | 
lemma bounded_by_update_var:  | 
| 60680 | 987  | 
assumes "bounded_by xs vs"  | 
988  | 
and "vs ! i = Some (l, u)"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
989  | 
    and bnd: "x \<in> { real_of_float l .. real_of_float u }"
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
990  | 
shows "bounded_by (xs[i := x]) vs"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
991  | 
proof (cases "i < length xs")  | 
| 49351 | 992  | 
case False  | 
| 60680 | 993  | 
thus ?thesis  | 
994  | 
using \<open>bounded_by xs vs\<close> by auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
995  | 
next  | 
| 60680 | 996  | 
case True  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
997  | 
let ?xs = "xs[i := x]"  | 
| 60680 | 998  | 
from True have "i < length ?xs" by auto  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
999  | 
  have "case vs ! j of None \<Rightarrow> True | Some (l, u) \<Rightarrow> ?xs ! j \<in> {real_of_float l .. real_of_float u}"
 | 
| 60680 | 1000  | 
if "j < length vs" for j  | 
1001  | 
proof (cases "vs ! j")  | 
|
1002  | 
case None  | 
|
1003  | 
then show ?thesis by simp  | 
|
1004  | 
next  | 
|
1005  | 
case (Some b)  | 
|
1006  | 
thus ?thesis  | 
|
1007  | 
proof (cases "i = j")  | 
|
1008  | 
case True  | 
|
1009  | 
thus ?thesis using \<open>vs ! i = Some (l, u)\<close> Some and bnd \<open>i < length ?xs\<close>  | 
|
1010  | 
by auto  | 
|
1011  | 
next  | 
|
1012  | 
case False  | 
|
| 49351 | 1013  | 
thus ?thesis  | 
| 60680 | 1014  | 
using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>j < length vs\<close>] Some by auto  | 
1015  | 
qed  | 
|
1016  | 
qed  | 
|
1017  | 
thus ?thesis  | 
|
1018  | 
unfolding bounded_by_def by auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1019  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1020  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1021  | 
lemma isDERIV_approx':  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1022  | 
assumes "bounded_by xs vs"  | 
| 60680 | 1023  | 
and vs_x: "vs ! x = Some (l, u)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1024  | 
    and X_in: "X \<in> {real_of_float l .. real_of_float u}"
 | 
| 49351 | 1025  | 
and approx: "isDERIV_approx prec x f vs"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1026  | 
shows "isDERIV x f (xs[x := X])"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1027  | 
proof -  | 
| 60680 | 1028  | 
from bounded_by_update_var[OF \<open>bounded_by xs vs\<close> vs_x X_in] approx  | 
1029  | 
show ?thesis by (rule isDERIV_approx)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1030  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1031  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1032  | 
lemma DERIV_approx:  | 
| 60680 | 1033  | 
assumes "n < length xs"  | 
1034  | 
and bnd: "bounded_by xs vs"  | 
|
| 49351 | 1035  | 
and isD: "isDERIV_approx prec n f vs"  | 
1036  | 
and app: "Some (l, u) = approx prec (DERIV_floatarith n f) vs" (is "_ = approx _ ?D _")  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1037  | 
shows "\<exists>(x::real). l \<le> x \<and> x \<le> u \<and>  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1038  | 
DERIV (\<lambda> x. interpret_floatarith f (xs[n := x])) (xs!n) :> x"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1039  | 
(is "\<exists> x. _ \<and> _ \<and> DERIV (?i f) _ :> _")  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1040  | 
proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI[OF _ conjI])  | 
| 60680 | 1041  | 
let "?i f" = "\<lambda>x. interpret_floatarith f (xs[n := x])"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1042  | 
from approx[OF bnd app]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1043  | 
show "l \<le> ?i ?D (xs!n)" and "?i ?D (xs!n) \<le> u"  | 
| 60533 | 1044  | 
using \<open>n < length xs\<close> by auto  | 
1045  | 
from DERIV_floatarith[OF \<open>n < length xs\<close>, of f "xs!n"] isDERIV_approx[OF bnd isD]  | 
|
| 60680 | 1046  | 
show "DERIV (?i f) (xs!n) :> (?i ?D (xs!n))"  | 
1047  | 
by simp  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1048  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1049  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
1050  | 
lemma lift_bin_aux:  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1051  | 
assumes lift_bin_Some: "Some (l, u) = lift_bin a b f"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1052  | 
obtains l1 u1 l2 u2  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1053  | 
where "a = Some (l1, u1)"  | 
| 49351 | 1054  | 
and "b = Some (l2, u2)"  | 
1055  | 
and "f l1 u1 l2 u2 = Some (l, u)"  | 
|
1056  | 
using assms by (cases a, simp, cases b, simp, auto)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1057  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
1058  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1059  | 
fun approx_tse where  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1060  | 
"approx_tse prec n 0 c k f bs = approx prec f bs" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1061  | 
"approx_tse prec n (Suc s) c k f bs =  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1062  | 
(if isDERIV_approx prec n f bs then  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1063  | 
lift_bin (approx prec f (bs[n := Some (c,c)]))  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1064  | 
(approx_tse prec n s c (Suc k) (DERIV_floatarith n f) bs)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1065  | 
(\<lambda> l1 u1 l2 u2. approx prec  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1066  | 
(Add (Var 0)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1067  | 
(Mult (Inverse (Num (Float (int k) 0)))  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1068  | 
(Mult (Add (Var (Suc (Suc 0))) (Minus (Num c)))  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1069  | 
(Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), bs!n])  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1070  | 
else approx prec f bs)"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1071  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1072  | 
lemma bounded_by_Cons:  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1073  | 
assumes bnd: "bounded_by xs vs"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1074  | 
    and x: "x \<in> { real_of_float l .. real_of_float u }"
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1075  | 
shows "bounded_by (x#xs) ((Some (l, u))#vs)"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1076  | 
proof -  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1077  | 
  have "case ((Some (l,u))#vs) ! i of Some (l, u) \<Rightarrow> (x#xs)!i \<in> { real_of_float l .. real_of_float u } | None \<Rightarrow> True"
 | 
| 60680 | 1078  | 
if *: "i < length ((Some (l, u))#vs)" for i  | 
1079  | 
proof (cases i)  | 
|
1080  | 
case 0  | 
|
1081  | 
with x show ?thesis by auto  | 
|
1082  | 
next  | 
|
1083  | 
case (Suc i)  | 
|
1084  | 
with * have "i < length vs" by auto  | 
|
1085  | 
from bnd[THEN bounded_byE, OF this]  | 
|
1086  | 
show ?thesis unfolding Suc nth_Cons_Suc .  | 
|
1087  | 
qed  | 
|
1088  | 
thus ?thesis  | 
|
1089  | 
by (auto simp add: bounded_by_def)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1090  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1091  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1092  | 
lemma approx_tse_generic:  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1093  | 
assumes "bounded_by xs vs"  | 
| 60680 | 1094  | 
and bnd_c: "bounded_by (xs[x := c]) vs"  | 
1095  | 
and "x < length vs" and "x < length xs"  | 
|
| 49351 | 1096  | 
and bnd_x: "vs ! x = Some (lx, ux)"  | 
1097  | 
and ate: "Some (l, u) = approx_tse prec x s c k f vs"  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1098  | 
  shows "\<exists> n. (\<forall> m < n. \<forall> (z::real) \<in> {lx .. ux}.
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1099  | 
DERIV (\<lambda> y. interpret_floatarith ((DERIV_floatarith x ^^ m) f) (xs[x := y])) z :>  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1100  | 
(interpret_floatarith ((DERIV_floatarith x ^^ (Suc m)) f) (xs[x := z])))  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1101  | 
   \<and> (\<forall> (t::real) \<in> {lx .. ux}.  (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) *
 | 
| 
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1102  | 
interpret_floatarith ((DERIV_floatarith x ^^ i) f) (xs[x := c]) *  | 
| 
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1103  | 
(xs!x - c)^i) +  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1104  | 
      inverse (real (\<Prod> j \<in> {k..<k+n}. j)) *
 | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1105  | 
interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := t]) *  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1106  | 
      (xs!x - c)^n \<in> {l .. u})" (is "\<exists> n. ?taylor f k l u n")
 | 
| 60680 | 1107  | 
using ate  | 
1108  | 
proof (induct s arbitrary: k f l u)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1109  | 
case 0  | 
| 49351 | 1110  | 
  {
 | 
1111  | 
    fix t::real assume "t \<in> {lx .. ux}"
 | 
|
| 60533 | 1112  | 
note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> bnd_x this]  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1113  | 
from approx[OF this 0[unfolded approx_tse.simps]]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1114  | 
    have "(interpret_floatarith f (xs[x := t])) \<in> {l .. u}"
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1115  | 
by (auto simp add: algebra_simps)  | 
| 49351 | 1116  | 
}  | 
1117  | 
thus ?case by (auto intro!: exI[of _ 0])  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1118  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1119  | 
case (Suc s)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1120  | 
show ?case  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1121  | 
proof (cases "isDERIV_approx prec x f vs")  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1122  | 
case False  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1123  | 
note ap = Suc.prems[unfolded approx_tse.simps if_not_P[OF False]]  | 
| 49351 | 1124  | 
    {
 | 
1125  | 
      fix t::real assume "t \<in> {lx .. ux}"
 | 
|
| 60533 | 1126  | 
note bounded_by_update_var[OF \<open>bounded_by xs vs\<close> bnd_x this]  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1127  | 
from approx[OF this ap]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1128  | 
      have "(interpret_floatarith f (xs[x := t])) \<in> {l .. u}"
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1129  | 
by (auto simp add: algebra_simps)  | 
| 49351 | 1130  | 
}  | 
1131  | 
thus ?thesis by (auto intro!: exI[of _ 0])  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1132  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1133  | 
case True  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1134  | 
with Suc.prems  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1135  | 
obtain l1 u1 l2 u2  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1136  | 
where a: "Some (l1, u1) = approx prec f (vs[x := Some (c,c)])"  | 
| 49351 | 1137  | 
and ate: "Some (l2, u2) = approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs"  | 
1138  | 
and final: "Some (l, u) = approx prec  | 
|
1139  | 
(Add (Var 0)  | 
|
1140  | 
(Mult (Inverse (Num (Float (int k) 0)))  | 
|
1141  | 
(Mult (Add (Var (Suc (Suc 0))) (Minus (Num c)))  | 
|
1142  | 
(Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), vs!x]"  | 
|
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
1143  | 
by (auto elim!: lift_bin_aux)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1144  | 
|
| 60533 | 1145  | 
from bnd_c \<open>x < length xs\<close>  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1146  | 
have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (c,c)])"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1147  | 
by (auto intro!: bounded_by_update)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1148  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1149  | 
from approx[OF this a]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1150  | 
    have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \<in> { l1 .. u1 }"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1151  | 
(is "?f 0 (real_of_float c) \<in> _")  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1152  | 
by auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1153  | 
|
| 60680 | 1154  | 
have funpow_Suc[symmetric]: "(f ^^ Suc n) x = (f ^^ n) (f x)"  | 
1155  | 
for f :: "'a \<Rightarrow> 'a" and n :: nat and x :: 'a  | 
|
1156  | 
by (induct n) auto  | 
|
1157  | 
from Suc.hyps[OF ate, unfolded this] obtain n  | 
|
1158  | 
      where DERIV_hyp: "\<And>m z. \<lbrakk> m < n ; (z::real) \<in> { lx .. ux } \<rbrakk> \<Longrightarrow>
 | 
|
1159  | 
DERIV (?f (Suc m)) z :> ?f (Suc (Suc m)) z"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1160  | 
      and hyp: "\<forall>t \<in> {real_of_float lx .. real_of_float ux}.
 | 
| 60680 | 1161  | 
        (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {Suc k..<Suc k + i}. j)) * ?f (Suc i) c * (xs!x - c)^i) +
 | 
1162  | 
          inverse (real (\<Prod> j \<in> {Suc k..<Suc k + n}. j)) * ?f (Suc n) t * (xs!x - c)^n \<in> {l2 .. u2}"
 | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1163  | 
(is "\<forall> t \<in> _. ?X (Suc k) f n t \<in> _")  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1164  | 
by blast  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1165  | 
|
| 60680 | 1166  | 
have DERIV: "DERIV (?f m) z :> ?f (Suc m) z"  | 
1167  | 
      if "m < Suc n" and bnd_z: "z \<in> { lx .. ux }" for m and z::real
 | 
|
1168  | 
proof (cases m)  | 
|
1169  | 
case 0  | 
|
1170  | 
with DERIV_floatarith[OF \<open>x < length xs\<close>  | 
|
1171  | 
isDERIV_approx'[OF \<open>bounded_by xs vs\<close> bnd_x bnd_z True]]  | 
|
1172  | 
show ?thesis by simp  | 
|
1173  | 
next  | 
|
1174  | 
case (Suc m')  | 
|
1175  | 
hence "m' < n"  | 
|
1176  | 
using \<open>m < Suc n\<close> by auto  | 
|
1177  | 
from DERIV_hyp[OF this bnd_z] show ?thesis  | 
|
1178  | 
using Suc by simp  | 
|
1179  | 
qed  | 
|
1180  | 
||
1181  | 
    have "\<And>k i. k < i \<Longrightarrow> {k ..< i} = insert k {Suc k ..< i}" by auto
 | 
|
| 64272 | 1182  | 
    hence prod_head_Suc: "\<And>k i. \<Prod>{k ..< k + Suc i} = k * \<Prod>{Suc k ..< Suc k + i}"
 | 
| 60680 | 1183  | 
by auto  | 
| 64267 | 1184  | 
    have sum_move0: "\<And>k F. sum F {0..<Suc k} = F 0 + sum (\<lambda> k. F (Suc k)) {0..<k}"
 | 
1185  | 
unfolding sum_shift_bounds_Suc_ivl[symmetric]  | 
|
1186  | 
unfolding sum_head_upt_Suc[OF zero_less_Suc] ..  | 
|
| 63040 | 1187  | 
define C where "C = xs!x - c"  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1188  | 
|
| 49351 | 1189  | 
    {
 | 
1190  | 
      fix t::real assume t: "t \<in> {lx .. ux}"
 | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1191  | 
hence "bounded_by [xs!x] [vs!x]"  | 
| 60533 | 1192  | 
using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>]  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1193  | 
by (cases "vs!x", auto simp add: bounded_by_def)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1194  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1195  | 
with hyp[THEN bspec, OF t] f_c  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1196  | 
have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some (l1, u1), Some (l2, u2), vs!x]"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1197  | 
by (auto intro!: bounded_by_Cons)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1198  | 
from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]]  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1199  | 
      have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \<in> {l .. u}"
 | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1200  | 
by (auto simp add: algebra_simps)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1201  | 
also have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse (real k) + ?f 0 c =  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1202  | 
               (\<Sum> i = 0..<Suc n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) * ?f i c * (xs!x - c)^i) +
 | 
| 
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1203  | 
               inverse (real (\<Prod> j \<in> {k..<k+Suc n}. j)) * ?f (Suc n) t * (xs!x - c)^Suc n" (is "_ = ?T")
 | 
| 64272 | 1204  | 
unfolding funpow_Suc C_def[symmetric] sum_move0 prod_head_Suc  | 
| 35082 | 1205  | 
by (auto simp add: algebra_simps)  | 
| 64267 | 1206  | 
(simp only: mult.left_commute [of _ "inverse (real k)"] sum_distrib_left [symmetric])  | 
| 49351 | 1207  | 
      finally have "?T \<in> {l .. u}" .
 | 
1208  | 
}  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1209  | 
thus ?thesis using DERIV by blast  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1210  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1211  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1212  | 
|
| 64272 | 1213  | 
lemma prod_fact: "real (\<Prod> {1..<1 + k}) = fact (k :: nat)"
 | 
1214  | 
by (simp add: fact_prod atLeastLessThanSuc_atLeastAtMost)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1215  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1216  | 
lemma approx_tse:  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1217  | 
assumes "bounded_by xs vs"  | 
| 60680 | 1218  | 
and bnd_x: "vs ! x = Some (lx, ux)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1219  | 
    and bnd_c: "real_of_float c \<in> {lx .. ux}"
 | 
| 49351 | 1220  | 
and "x < length vs" and "x < length xs"  | 
1221  | 
and ate: "Some (l, u) = approx_tse prec x s c 1 f vs"  | 
|
| 60680 | 1222  | 
  shows "interpret_floatarith f xs \<in> {l .. u}"
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1223  | 
proof -  | 
| 63040 | 1224  | 
define F where [abs_def]: "F n z =  | 
1225  | 
interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := z])" for n z  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1226  | 
hence F0: "F 0 = (\<lambda> z. interpret_floatarith f (xs[x := z]))" by auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1227  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1228  | 
hence "bounded_by (xs[x := c]) vs" and "x < length vs" "x < length xs"  | 
| 60533 | 1229  | 
using \<open>bounded_by xs vs\<close> bnd_x bnd_c \<open>x < length vs\<close> \<open>x < length xs\<close>  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1230  | 
by (auto intro!: bounded_by_update_var)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1231  | 
|
| 60533 | 1232  | 
from approx_tse_generic[OF \<open>bounded_by xs vs\<close> this bnd_x ate]  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1233  | 
obtain n  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1234  | 
where DERIV: "\<forall> m z. m < n \<and> real_of_float lx \<le> z \<and> z \<le> real_of_float ux \<longrightarrow> DERIV (F m) z :> F (Suc m) z"  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1235  | 
    and hyp: "\<And> (t::real). t \<in> {lx .. ux} \<Longrightarrow>
 | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59658 
diff
changeset
 | 
1236  | 
(\<Sum> j = 0..<n. inverse(fact j) * F j c * (xs!x - c)^j) +  | 
| 
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59658 
diff
changeset
 | 
1237  | 
inverse ((fact n)) * F n t * (xs!x - c)^n  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1238  | 
             \<in> {l .. u}" (is "\<And> t. _ \<Longrightarrow> ?taylor t \<in> _")
 | 
| 64272 | 1239  | 
unfolding F_def atLeastAtMost_iff[symmetric] prod_fact  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59658 
diff
changeset
 | 
1240  | 
by blast  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1241  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1242  | 
  have bnd_xs: "xs ! x \<in> { lx .. ux }"
 | 
| 60533 | 1243  | 
using \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1244  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1245  | 
show ?thesis  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1246  | 
proof (cases n)  | 
| 60680 | 1247  | 
case 0  | 
1248  | 
thus ?thesis  | 
|
1249  | 
using hyp[OF bnd_xs] unfolding F_def by auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1250  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1251  | 
case (Suc n')  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1252  | 
show ?thesis  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1253  | 
proof (cases "xs ! x = c")  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1254  | 
case True  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1255  | 
from True[symmetric] hyp[OF bnd_xs] Suc show ?thesis  | 
| 64267 | 1256  | 
unfolding F_def Suc sum_head_upt_Suc[OF zero_less_Suc] sum_shift_bounds_Suc_ivl  | 
| 60680 | 1257  | 
by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1258  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1259  | 
case False  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1260  | 
have "lx \<le> real_of_float c" "real_of_float c \<le> ux" "lx \<le> xs!x" "xs!x \<le> ux"  | 
| 60533 | 1261  | 
using Suc bnd_c \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x by auto  | 
| 63570 | 1262  | 
from taylor[OF zero_less_Suc, of F, OF F0 DERIV[unfolded Suc] this False]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1263  | 
obtain t::real where t_bnd: "if xs ! x < c then xs ! x < t \<and> t < c else c < t \<and> t < xs ! x"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1264  | 
and fl_eq: "interpret_floatarith f (xs[x := xs ! x]) =  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59658 
diff
changeset
 | 
1265  | 
(\<Sum>m = 0..<Suc n'. F m c / (fact m) * (xs ! x - c) ^ m) +  | 
| 
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59658 
diff
changeset
 | 
1266  | 
F (Suc n') t / (fact (Suc n')) * (xs ! x - c) ^ Suc n'"  | 
| 56195 | 1267  | 
unfolding atLeast0LessThan by blast  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1268  | 
|
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1269  | 
      from t_bnd bnd_xs bnd_c have *: "t \<in> {lx .. ux}"
 | 
| 60680 | 1270  | 
by (cases "xs ! x < c") auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1271  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1272  | 
have "interpret_floatarith f (xs[x := xs ! x]) = ?taylor t"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1273  | 
unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse)  | 
| 60680 | 1274  | 
      also have "\<dots> \<in> {l .. u}"
 | 
1275  | 
using * by (rule hyp)  | 
|
1276  | 
finally show ?thesis  | 
|
1277  | 
by simp  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1278  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1279  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1280  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1281  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1282  | 
fun approx_tse_form' where  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1283  | 
"approx_tse_form' prec t f 0 l u cmp =  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1284  | 
(case approx_tse prec 0 t ((l + u) * Float 1 (- 1)) 1 f [Some (l, u)]  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1285  | 
of Some (l, u) \<Rightarrow> cmp l u | None \<Rightarrow> False)" |  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1286  | 
"approx_tse_form' prec t f (Suc s) l u cmp =  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1287  | 
(let m = (l + u) * Float 1 (- 1)  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1288  | 
in (if approx_tse_form' prec t f s l m cmp then  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1289  | 
approx_tse_form' prec t f s m u cmp else False))"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1290  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1291  | 
lemma approx_tse_form':  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1292  | 
fixes x :: real  | 
| 60680 | 1293  | 
assumes "approx_tse_form' prec t f s l u cmp"  | 
1294  | 
    and "x \<in> {l .. u}"
 | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1295  | 
  shows "\<exists>l' u' ly uy. x \<in> {l' .. u'} \<and> real_of_float l \<le> l' \<and> u' \<le> real_of_float u \<and> cmp ly uy \<and>
 | 
| 60680 | 1296  | 
approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)"  | 
1297  | 
using assms  | 
|
1298  | 
proof (induct s arbitrary: l u)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1299  | 
case 0  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1300  | 
then obtain ly uy  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1301  | 
where *: "approx_tse prec 0 t ((l + u) * Float 1 (- 1)) 1 f [Some (l, u)] = Some (ly, uy)"  | 
| 
55413
 
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
 
blanchet 
parents: 
54782 
diff
changeset
 | 
1302  | 
and **: "cmp ly uy" by (auto elim!: case_optionE)  | 
| 46545 | 1303  | 
with 0 show ?case by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1304  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1305  | 
case (Suc s)  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1306  | 
let ?m = "(l + u) * Float 1 (- 1)"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1307  | 
from Suc.prems  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1308  | 
have l: "approx_tse_form' prec t f s l ?m cmp"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1309  | 
and u: "approx_tse_form' prec t f s ?m u cmp"  | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1310  | 
by (auto simp add: Let_def lazy_conj)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1311  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1312  | 
have m_l: "real_of_float l \<le> ?m" and m_u: "?m \<le> real_of_float u"  | 
| 
47599
 
400b158f1589
replace the float datatype by a type with unique representation
 
hoelzl 
parents: 
47108 
diff
changeset
 | 
1313  | 
unfolding less_eq_float_def using Suc.prems by auto  | 
| 60680 | 1314  | 
  with \<open>x \<in> { l .. u }\<close> consider "x \<in> { l .. ?m}" | "x \<in> {?m .. u}"
 | 
1315  | 
by atomize_elim auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1316  | 
thus ?case  | 
| 60680 | 1317  | 
proof cases  | 
1318  | 
case 1  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1319  | 
from Suc.hyps[OF l this]  | 
| 60680 | 1320  | 
obtain l' u' ly uy where  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1321  | 
      "x \<in> {l' .. u'} \<and> real_of_float l \<le> l' \<and> real_of_float u' \<le> ?m \<and> cmp ly uy \<and>
 | 
| 60680 | 1322  | 
approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)"  | 
1323  | 
by blast  | 
|
1324  | 
with m_u show ?thesis  | 
|
1325  | 
by (auto intro!: exI)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1326  | 
next  | 
| 60680 | 1327  | 
case 2  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1328  | 
from Suc.hyps[OF u this]  | 
| 60680 | 1329  | 
obtain l' u' ly uy where  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1330  | 
      "x \<in> { l' .. u' } \<and> ?m \<le> real_of_float l' \<and> u' \<le> real_of_float u \<and> cmp ly uy \<and>
 | 
| 60680 | 1331  | 
approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 f [Some (l', u')] = Some (ly, uy)"  | 
1332  | 
by blast  | 
|
1333  | 
with m_u show ?thesis  | 
|
1334  | 
by (auto intro!: exI)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1335  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1336  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1337  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1338  | 
lemma approx_tse_form'_less:  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1339  | 
fixes x :: real  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1340  | 
assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\<lambda> l u. 0 < l)"  | 
| 60680 | 1341  | 
    and x: "x \<in> {l .. u}"
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1342  | 
shows "interpret_floatarith b [x] < interpret_floatarith a [x]"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1343  | 
proof -  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1344  | 
from approx_tse_form'[OF tse x]  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1345  | 
obtain l' u' ly uy  | 
| 60680 | 1346  | 
    where x': "x \<in> {l' .. u'}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1347  | 
and "real_of_float l \<le> real_of_float l'"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1348  | 
and "real_of_float u' \<le> real_of_float u" and "0 < ly"  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1349  | 
and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1350  | 
by blast  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1351  | 
|
| 60680 | 1352  | 
hence "bounded_by [x] [Some (l', u')]"  | 
1353  | 
by (auto simp add: bounded_by_def)  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1354  | 
from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1355  | 
have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
53077 
diff
changeset
 | 
1356  | 
by auto  | 
| 60680 | 1357  | 
from order_less_le_trans[OF _ this, of 0] \<open>0 < ly\<close> show ?thesis  | 
1358  | 
by auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1359  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1360  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1361  | 
lemma approx_tse_form'_le:  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1362  | 
fixes x :: real  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1363  | 
assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s l u (\<lambda> l u. 0 \<le> l)"  | 
| 60680 | 1364  | 
    and x: "x \<in> {l .. u}"
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1365  | 
shows "interpret_floatarith b [x] \<le> interpret_floatarith a [x]"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1366  | 
proof -  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1367  | 
from approx_tse_form'[OF tse x]  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1368  | 
obtain l' u' ly uy  | 
| 60680 | 1369  | 
    where x': "x \<in> {l' .. u'}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1370  | 
and "l \<le> real_of_float l'"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60680 
diff
changeset
 | 
1371  | 
and "real_of_float u' \<le> u" and "0 \<le> ly"  | 
| 
58410
 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 
haftmann 
parents: 
58310 
diff
changeset
 | 
1372  | 
and tse: "approx_tse prec 0 t ((l' + u') * Float 1 (- 1)) 1 (Add a (Minus b)) [Some (l', u')] = Some (ly, uy)"  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1373  | 
by blast  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1374  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1375  | 
hence "bounded_by [x] [Some (l', u')]" by (auto simp add: bounded_by_def)  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1376  | 
from approx_tse[OF this _ _ _ _ tse[symmetric], of l' u'] x'  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1377  | 
have "ly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]"  | 
| 
54230
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
53077 
diff
changeset
 | 
1378  | 
by auto  | 
| 60680 | 1379  | 
from order_trans[OF _ this, of 0] \<open>0 \<le> ly\<close> show ?thesis  | 
1380  | 
by auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1381  | 
qed  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1382  | 
|
| 58986 | 1383  | 
fun approx_tse_concl where  | 
1384  | 
"approx_tse_concl prec t (Less lf rt) s l u l' u' \<longleftrightarrow>  | 
|
1385  | 
approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)" |  | 
|
1386  | 
"approx_tse_concl prec t (LessEqual lf rt) s l u l' u' \<longleftrightarrow>  | 
|
1387  | 
approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)" |  | 
|
1388  | 
"approx_tse_concl prec t (AtLeastAtMost x lf rt) s l u l' u' \<longleftrightarrow>  | 
|
1389  | 
(if approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l) then  | 
|
1390  | 
approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l) else False)" |  | 
|
1391  | 
"approx_tse_concl prec t (Conj f g) s l u l' u' \<longleftrightarrow>  | 
|
1392  | 
approx_tse_concl prec t f s l u l' u' \<and> approx_tse_concl prec t g s l u l' u'" |  | 
|
1393  | 
"approx_tse_concl prec t (Disj f g) s l u l' u' \<longleftrightarrow>  | 
|
1394  | 
approx_tse_concl prec t f s l u l' u' \<or> approx_tse_concl prec t g s l u l' u'" |  | 
|
1395  | 
"approx_tse_concl _ _ _ _ _ _ _ _ \<longleftrightarrow> False"  | 
|
1396  | 
||
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1397  | 
definition  | 
| 60680 | 1398  | 
"approx_tse_form prec t s f =  | 
1399  | 
(case f of  | 
|
1400  | 
Bound x a b f \<Rightarrow>  | 
|
1401  | 
x = Var 0 \<and>  | 
|
1402  | 
(case (approx prec a [None], approx prec b [None]) of  | 
|
1403  | 
(Some (l, u), Some (l', u')) \<Rightarrow> approx_tse_concl prec t f s l u l' u'  | 
|
1404  | 
| _ \<Rightarrow> False)  | 
|
1405  | 
| _ \<Rightarrow> False)"  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1406  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1407  | 
lemma approx_tse_form:  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1408  | 
assumes "approx_tse_form prec t s f"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1409  | 
shows "interpret_form f [x]"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1410  | 
proof (cases f)  | 
| 60680 | 1411  | 
case f_def: (Bound i a b f')  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1412  | 
with assms obtain l u l' u'  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1413  | 
where a: "approx prec a [None] = Some (l, u)"  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1414  | 
and b: "approx prec b [None] = Some (l', u')"  | 
| 
55413
 
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
 
blanchet 
parents: 
54782 
diff
changeset
 | 
1415  | 
unfolding approx_tse_form_def by (auto elim!: case_optionE)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1416  | 
|
| 60680 | 1417  | 
from f_def assms have "i = Var 0"  | 
1418  | 
unfolding approx_tse_form_def by auto  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1419  | 
hence i: "interpret_floatarith i [x] = x" by auto  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1420  | 
|
| 60680 | 1421  | 
  {
 | 
1422  | 
let ?f = "\<lambda>z. interpret_floatarith z [x]"  | 
|
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1423  | 
    assume "?f i \<in> { ?f a .. ?f b }"
 | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1424  | 
with approx[OF _ a[symmetric], of "[x]"] approx[OF _ b[symmetric], of "[x]"]  | 
| 
40881
 
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
 
hoelzl 
parents: 
39556 
diff
changeset
 | 
1425  | 
    have bnd: "x \<in> { l .. u'}" unfolding bounded_by_def i by auto
 | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1426  | 
|
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1427  | 
have "interpret_form f' [x]"  | 
| 60680 | 1428  | 
using assms[unfolded f_def]  | 
| 58986 | 1429  | 
proof (induct f')  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1430  | 
case (Less lf rt)  | 
| 58986 | 1431  | 
with a b  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1432  | 
have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 < l)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1433  | 
unfolding approx_tse_form_def by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1434  | 
from approx_tse_form'_less[OF this bnd]  | 
| 58986 | 1435  | 
show ?case using Less by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1436  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1437  | 
case (LessEqual lf rt)  | 
| 60680 | 1438  | 
with f_def a b assms  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1439  | 
have "approx_tse_form' prec t (Add rt (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1440  | 
unfolding approx_tse_form_def by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1441  | 
from approx_tse_form'_le[OF this bnd]  | 
| 58986 | 1442  | 
show ?case using LessEqual by auto  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1443  | 
next  | 
| 
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1444  | 
case (AtLeastAtMost x lf rt)  | 
| 60680 | 1445  | 
with f_def a b assms  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1446  | 
have "approx_tse_form' prec t (Add rt (Minus x)) s l u' (\<lambda> l u. 0 \<le> l)"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
32920 
diff
changeset
 | 
1447  | 
and "approx_tse_form' prec t (Add x (Minus lf)) s l u' (\<lambda> l u. 0 \<le> l)"  | 
| 62390 | 1448  | 
unfolding approx_tse_form_def lazy_conj by (auto split: if_split_asm)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1449  | 
from approx_tse_form'_le[OF this(1) bnd] approx_tse_form'_le[OF this(2) bnd]  | 
| 58986 | 1450  | 
show ?case using AtLeastAtMost by auto  | 
1451  | 
qed (auto simp: f_def approx_tse_form_def elim!: case_optionE)  | 
|
| 60680 | 1452  | 
}  | 
1453  | 
thus ?thesis unfolding f_def by auto  | 
|
| 58986 | 1454  | 
qed (insert assms, auto simp add: approx_tse_form_def)  | 
| 
31863
 
e391eee8bf14
Implemented taylor series expansion for approximation
 
hoelzl 
parents: 
31811 
diff
changeset
 | 
1455  | 
|
| 60533 | 1456  | 
text \<open>@{term approx_form_eval} is only used for the {\tt value}-command.\<close>
 | 
| 
32919
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1457  | 
|
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1458  | 
fun approx_form_eval :: "nat \<Rightarrow> form \<Rightarrow> (float * float) option list \<Rightarrow> (float * float) option list" where  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1459  | 
"approx_form_eval prec (Bound (Var n) a b f) bs =  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1460  | 
(case (approx prec a bs, approx prec b bs)  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1461  | 
of (Some (l, _), Some (_, u)) \<Rightarrow> approx_form_eval prec f (bs[n := Some (l, u)])  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1462  | 
| _ \<Rightarrow> bs)" |  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1463  | 
"approx_form_eval prec (Assign (Var n) a f) bs =  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1464  | 
(case (approx prec a bs)  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1465  | 
of (Some (l, u)) \<Rightarrow> approx_form_eval prec f (bs[n := Some (l, u)])  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1466  | 
| _ \<Rightarrow> bs)" |  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1467  | 
"approx_form_eval prec (Less a b) bs = bs @ [approx prec a bs, approx prec b bs]" |  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1468  | 
"approx_form_eval prec (LessEqual a b) bs = bs @ [approx prec a bs, approx prec b bs]" |  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1469  | 
"approx_form_eval prec (AtLeastAtMost x a b) bs =  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1470  | 
bs @ [approx prec x bs, approx prec a bs, approx prec b bs]" |  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1471  | 
"approx_form_eval _ _ bs = bs"  | 
| 
 
37adfa07b54b
approximation now fails earlier when using interval splitting; value [approximate] now supports bounded variables; renamed Var -> Atom for better readability
 
hoelzl 
parents: 
32650 
diff
changeset
 | 
1472  | 
|
| 60680 | 1473  | 
|
| 60533 | 1474  | 
subsection \<open>Implement proof method \texttt{approximation}\<close>
 | 
| 29805 | 1475  | 
|
| 60533 | 1476  | 
oracle approximation_oracle = \<open>fn (thy, t) =>  | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1477  | 
let  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1478  | 
  fun bad t = error ("Bad term: " ^ Syntax.string_of_term_global thy t);
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1479  | 
|
| 
38716
 
3c3b4ad683d5
approximation_oracle: actually match true/false in ML, not arbitrary values;
 
wenzelm 
parents: 
38558 
diff
changeset
 | 
1480  | 
  fun term_of_bool true = @{term True}
 | 
| 
 
3c3b4ad683d5
approximation_oracle: actually match true/false in ML, not arbitrary values;
 
wenzelm 
parents: 
38558 
diff
changeset
 | 
1481  | 
    | term_of_bool false = @{term False};
 | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1482  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1483  | 
  val mk_int = HOLogic.mk_number @{typ int} o @{code integer_of_int};
 | 
| 58988 | 1484  | 
  fun dest_int (@{term int_of_integer} $ j) = @{code int_of_integer} (snd (HOLogic.dest_number j))
 | 
1485  | 
    | dest_int i = @{code int_of_integer} (snd (HOLogic.dest_number i));
 | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1486  | 
|
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1487  | 
  fun term_of_float (@{code Float} (k, l)) =
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1488  | 
    @{term Float} $ mk_int k $ mk_int l;
 | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1489  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1490  | 
  fun term_of_float_float_option NONE = @{term "None :: (float \<times> float) option"}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1491  | 
    | term_of_float_float_option (SOME ff) = @{term "Some :: float \<times> float \<Rightarrow> _"}
 | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58988 
diff
changeset
 | 
1492  | 
$ HOLogic.mk_prod (apply2 term_of_float ff);  | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1493  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1494  | 
val term_of_float_float_option_list =  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1495  | 
    HOLogic.mk_list @{typ "(float \<times> float) option"} o map term_of_float_float_option;
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1496  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1497  | 
  fun nat_of_term t = @{code nat_of_integer}
 | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1498  | 
(HOLogic.dest_nat t handle TERM _ => snd (HOLogic.dest_number t));  | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1499  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1500  | 
  fun float_of_term (@{term Float} $ k $ l) =
 | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
49962 
diff
changeset
 | 
1501  | 
        @{code Float} (dest_int k, dest_int l)
 | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1502  | 
| float_of_term t = bad t;  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1503  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1504  | 
  fun floatarith_of_term (@{term Add} $ a $ b) = @{code Add} (floatarith_of_term a, floatarith_of_term b)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1505  | 
    | floatarith_of_term (@{term Minus} $ a) = @{code Minus} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1506  | 
    | floatarith_of_term (@{term Mult} $ a $ b) = @{code Mult} (floatarith_of_term a, floatarith_of_term b)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1507  | 
    | floatarith_of_term (@{term Inverse} $ a) = @{code Inverse} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1508  | 
    | floatarith_of_term (@{term Cos} $ a) = @{code Cos} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1509  | 
    | floatarith_of_term (@{term Arctan} $ a) = @{code Arctan} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1510  | 
    | floatarith_of_term (@{term Abs} $ a) = @{code Abs} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1511  | 
    | floatarith_of_term (@{term Max} $ a $ b) = @{code Max} (floatarith_of_term a, floatarith_of_term b)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1512  | 
    | floatarith_of_term (@{term Min} $ a $ b) = @{code Min} (floatarith_of_term a, floatarith_of_term b)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1513  | 
    | floatarith_of_term @{term Pi} = @{code Pi}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1514  | 
    | floatarith_of_term (@{term Sqrt} $ a) = @{code Sqrt} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1515  | 
    | floatarith_of_term (@{term Exp} $ a) = @{code Exp} (floatarith_of_term a)
 | 
| 
62200
 
67792e4a5486
Made Approximation work for powr again
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
61969 
diff
changeset
 | 
1516  | 
    | floatarith_of_term (@{term Powr} $ a $ b) = @{code Powr} (floatarith_of_term a, floatarith_of_term b)
 | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1517  | 
    | floatarith_of_term (@{term Ln} $ a) = @{code Ln} (floatarith_of_term a)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1518  | 
    | floatarith_of_term (@{term Power} $ a $ n) =
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1519  | 
        @{code Power} (floatarith_of_term a, nat_of_term n)
 | 
| 
63263
 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 
immler 
parents: 
63248 
diff
changeset
 | 
1520  | 
    | floatarith_of_term (@{term Floor} $ a) = @{code Floor} (floatarith_of_term a)
 | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1521  | 
    | floatarith_of_term (@{term Var} $ n) = @{code Var} (nat_of_term n)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1522  | 
    | floatarith_of_term (@{term Num} $ m) = @{code Num} (float_of_term m)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1523  | 
| floatarith_of_term t = bad t;  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1524  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1525  | 
  fun form_of_term (@{term Bound} $ a $ b $ c $ p) = @{code Bound}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1526  | 
(floatarith_of_term a, floatarith_of_term b, floatarith_of_term c, form_of_term p)  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1527  | 
    | form_of_term (@{term Assign} $ a $ b $ p) = @{code Assign}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1528  | 
(floatarith_of_term a, floatarith_of_term b, form_of_term p)  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1529  | 
    | form_of_term (@{term Less} $ a $ b) = @{code Less}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1530  | 
(floatarith_of_term a, floatarith_of_term b)  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1531  | 
    | form_of_term (@{term LessEqual} $ a $ b) = @{code LessEqual}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1532  | 
(floatarith_of_term a, floatarith_of_term b)  | 
| 58986 | 1533  | 
    | form_of_term (@{term Conj} $ a $ b) = @{code Conj}
 | 
1534  | 
(form_of_term a, form_of_term b)  | 
|
1535  | 
    | form_of_term (@{term Disj} $ a $ b) = @{code Disj}
 | 
|
1536  | 
(form_of_term a, form_of_term b)  | 
|
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1537  | 
    | form_of_term (@{term AtLeastAtMost} $ a $ b $ c) = @{code AtLeastAtMost}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1538  | 
(floatarith_of_term a, floatarith_of_term b, floatarith_of_term c)  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1539  | 
| form_of_term t = bad t;  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1540  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1541  | 
  fun float_float_option_of_term @{term "None :: (float \<times> float) option"} = NONE
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1542  | 
    | float_float_option_of_term (@{term "Some :: float \<times> float \<Rightarrow> _"} $ ff) =
 | 
| 
59058
 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 
wenzelm 
parents: 
58988 
diff
changeset
 | 
1543  | 
SOME (apply2 float_of_term (HOLogic.dest_prod ff))  | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1544  | 
    | float_float_option_of_term (@{term approx'} $ n $ a $ ffs) = @{code approx'}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1545  | 
(nat_of_term n) (floatarith_of_term a) (float_float_option_list_of_term ffs)  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1546  | 
| float_float_option_of_term t = bad t  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1547  | 
and float_float_option_list_of_term  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1548  | 
        (@{term "replicate :: _ \<Rightarrow> (float \<times> float) option \<Rightarrow> _"} $ n $ @{term "None :: (float \<times> float) option"}) =
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1549  | 
          @{code replicate} (nat_of_term n) NONE
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1550  | 
    | float_float_option_list_of_term (@{term approx_form_eval} $ n $ p $ ffs) =
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1551  | 
        @{code approx_form_eval} (nat_of_term n) (form_of_term p) (float_float_option_list_of_term ffs)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1552  | 
| float_float_option_list_of_term t = map float_float_option_of_term  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1553  | 
(HOLogic.dest_list t);  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1554  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1555  | 
val nat_list_of_term = map nat_of_term o HOLogic.dest_list ;  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1556  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1557  | 
  fun bool_of_term (@{term approx_form} $ n $ p $ ffs $ ms) = @{code approx_form}
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1558  | 
(nat_of_term n) (form_of_term p) (float_float_option_list_of_term ffs) (nat_list_of_term ms)  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1559  | 
    | bool_of_term (@{term approx_tse_form} $ m $ n $ q $ p) =
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1560  | 
        @{code approx_tse_form} (nat_of_term m) (nat_of_term n) (nat_of_term q) (form_of_term p)
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1561  | 
| bool_of_term t = bad t;  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1562  | 
|
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1563  | 
fun eval t = case fastype_of t  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1564  | 
   of @{typ bool} =>
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1565  | 
(term_of_bool o bool_of_term) t  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1566  | 
    | @{typ "(float \<times> float) option"} =>
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1567  | 
(term_of_float_float_option o float_float_option_of_term) t  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1568  | 
    | @{typ "(float \<times> float) option list"} =>
 | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1569  | 
(term_of_float_float_option_list o float_float_option_list_of_term) t  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1570  | 
| _ => bad t;  | 
| 
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1571  | 
|
| 52131 | 1572  | 
val normalize = eval o Envir.beta_norm o Envir.eta_long [];  | 
| 
36985
 
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
 
haftmann 
parents: 
36960 
diff
changeset
 | 
1573  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
1574  | 
in Thm.global_cterm_of thy (Logic.mk_equals (t, normalize t)) end  | 
| 60533 | 1575  | 
\<close>  | 
| 
31099
 
03314c427b34
optimized Approximation by precompiling approx_inequality
 
hoelzl 
parents: 
31098 
diff
changeset
 | 
1576  | 
|
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1577  | 
lemma intervalE: "a \<le> x \<and> x \<le> b \<Longrightarrow> \<lbrakk> x \<in> { a .. b } \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1578  | 
by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1579  | 
|
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1580  | 
lemma meta_eqE: "x \<equiv> a \<Longrightarrow> \<lbrakk> x = a \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1581  | 
by auto  | 
| 
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1582  | 
|
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1583  | 
named_theorems approximation_preproc  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1584  | 
|
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1585  | 
lemma approximation_preproc_floatarith[approximation_preproc]:  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1586  | 
"0 = real_of_float 0"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1587  | 
"1 = real_of_float 1"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1588  | 
"0 = Float 0 0"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1589  | 
"1 = Float 1 0"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1590  | 
"numeral a = Float (numeral a) 0"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1591  | 
"numeral a = real_of_float (numeral a)"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1592  | 
"x - y = x + - y"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1593  | 
"x / y = x * inverse y"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1594  | 
"ceiling x = - floor (- x)"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1595  | 
"log x y = ln y * inverse (ln x)"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1596  | 
"sin x = cos (pi / 2 - x)"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1597  | 
"tan x = sin x / cos x"  | 
| 
63931
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1598  | 
by (simp_all add: inverse_eq_divide ceiling_def log_def sin_cos_eq tan_def real_of_float_eq)  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1599  | 
|
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1600  | 
lemma approximation_preproc_int[approximation_preproc]:  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1601  | 
"real_of_int 0 = real_of_float 0"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1602  | 
"real_of_int 1 = real_of_float 1"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1603  | 
"real_of_int (i + j) = real_of_int i + real_of_int j"  | 
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1604  | 
"real_of_int (- i) = - real_of_int i"  | 
| 
63931
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1605  | 
"real_of_int (i - j) = real_of_int i - real_of_int j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1606  | 
"real_of_int (i * j) = real_of_int i * real_of_int j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1607  | 
"real_of_int (i div j) = real_of_int (floor (real_of_int i / real_of_int j))"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1608  | 
"real_of_int (min i j) = min (real_of_int i) (real_of_int j)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1609  | 
"real_of_int (max i j) = max (real_of_int i) (real_of_int j)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1610  | 
"real_of_int (abs i) = abs (real_of_int i)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1611  | 
"real_of_int (i ^ n) = (real_of_int i) ^ n"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1612  | 
"real_of_int (numeral a) = real_of_float (numeral a)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1613  | 
"i mod j = i - i div j * j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1614  | 
"i = j \<longleftrightarrow> real_of_int i = real_of_int j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1615  | 
"i \<le> j \<longleftrightarrow> real_of_int i \<le> real_of_int j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1616  | 
"i < j \<longleftrightarrow> real_of_int i < real_of_int j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1617  | 
  "i \<in> {j .. k} \<longleftrightarrow> real_of_int i \<in> {real_of_int j .. real_of_int k}"
 | 
| 64246 | 1618  | 
by (simp_all add: floor_divide_of_int_eq minus_div_mult_eq_mod [symmetric])  | 
| 
63931
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1619  | 
|
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1620  | 
lemma approximation_preproc_nat[approximation_preproc]:  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1621  | 
"real 0 = real_of_float 0"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1622  | 
"real 1 = real_of_float 1"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1623  | 
"real (i + j) = real i + real j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1624  | 
"real (i - j) = max (real i - real j) 0"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1625  | 
"real (i * j) = real i * real j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1626  | 
"real (i div j) = real_of_int (floor (real i / real j))"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1627  | 
"real (min i j) = min (real i) (real j)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1628  | 
"real (max i j) = max (real i) (real j)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1629  | 
"real (i ^ n) = (real i) ^ n"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1630  | 
"real (numeral a) = real_of_float (numeral a)"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1631  | 
"i mod j = i - i div j * j"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1632  | 
"n = m \<longleftrightarrow> real n = real m"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1633  | 
"n \<le> m \<longleftrightarrow> real n \<le> real m"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1634  | 
"n < m \<longleftrightarrow> real n < real m"  | 
| 
 
f17a1c60ac39
approximation: preprocessing for nat/int expressions
 
immler 
parents: 
63929 
diff
changeset
 | 
1635  | 
  "n \<in> {m .. l} \<longleftrightarrow> real n \<in> {real m .. real l}"
 | 
| 64243 | 1636  | 
by (simp_all add: real_div_nat_eq_floor_of_divide minus_div_mult_eq_mod [symmetric])  | 
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1637  | 
|
| 59850 | 1638  | 
ML_file "approximation.ML"  | 
1639  | 
||
| 60533 | 1640  | 
method_setup approximation = \<open>  | 
| 60680 | 1641  | 
let  | 
1642  | 
val free =  | 
|
1643  | 
Args.context -- Args.term >> (fn (_, Free (n, _)) => n | (ctxt, t) =>  | 
|
1644  | 
        error ("Bad free variable: " ^ Syntax.string_of_term ctxt t));
 | 
|
| 59850 | 1645  | 
in  | 
| 60680 | 1646  | 
Scan.lift Parse.nat --  | 
| 59850 | 1647  | 
Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon)  | 
| 60680 | 1648  | 
|-- Parse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift Parse.nat)) [] --  | 
1649  | 
Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon) |--  | 
|
1650  | 
(free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift Parse.nat)) >>  | 
|
| 59850 | 1651  | 
(fn ((prec, splitting), taylor) => fn ctxt =>  | 
1652  | 
SIMPLE_METHOD' (Approximation.approximation_tac prec splitting taylor ctxt))  | 
|
1653  | 
end  | 
|
| 60533 | 1654  | 
\<close> "real number approximation"  | 
| 
31811
 
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
 
hoelzl 
parents: 
31810 
diff
changeset
 | 
1655  | 
|
| 58988 | 1656  | 
|
1657  | 
section "Quickcheck Generator"  | 
|
1658  | 
||
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1659  | 
lemma approximation_preproc_push_neg[approximation_preproc]:  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1660  | 
fixes a b::real  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1661  | 
shows  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1662  | 
"\<not> (a < b) \<longleftrightarrow> b \<le> a"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1663  | 
"\<not> (a \<le> b) \<longleftrightarrow> b < a"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1664  | 
"\<not> (a = b) \<longleftrightarrow> b < a \<or> a < b"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1665  | 
"\<not> (p \<and> q) \<longleftrightarrow> \<not> p \<or> \<not> q"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1666  | 
"\<not> (p \<or> q) \<longleftrightarrow> \<not> p \<and> \<not> q"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1667  | 
"\<not> \<not> q \<longleftrightarrow> q"  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1668  | 
by auto  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
63918 
diff
changeset
 | 
1669  | 
|
| 58988 | 1670  | 
ML_file "approximation_generator.ML"  | 
1671  | 
setup "Approximation_Generator.setup"  | 
|
1672  | 
||
| 29805 | 1673  | 
end  |