author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 74397 | e80c4cde6064 |
child 80034 | 95b4fb2b5359 |
permissions | -rw-r--r-- |
73546
7cb3fefef79e
confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
73537
diff
changeset
|
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 |
65582
a1bc1b020cf2
tuned Approximation: separated general material from oracle
eberlm <eberlm@in.tum.de>
parents:
65578
diff
changeset
|
7 |
Approximation_Bounds |
73546
7cb3fefef79e
confluent preprocessing for floats in presence of target language numerals
haftmann
parents:
73537
diff
changeset
|
8 |
"HOL-Library.Code_Target_Numeral_Float" |
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 |
||
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
76 |
fun lift_bin :: "(float interval) option \<Rightarrow> (float interval) option \<Rightarrow> (float interval \<Rightarrow> float interval \<Rightarrow> (float interval) option) \<Rightarrow> (float interval) option" where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
77 |
"lift_bin (Some ivl1) (Some ivl2) f = f ivl1 ivl2" | |
62200
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 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
80 |
fun lift_bin' :: "(float interval) option \<Rightarrow> (float interval) option \<Rightarrow> (float interval \<Rightarrow> float interval \<Rightarrow> float interval) \<Rightarrow> (float interval) option" where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
81 |
"lift_bin' (Some ivl1) (Some ivl2) f = Some (f ivl1 ivl2)" | |
29805 | 82 |
"lift_bin' a b f = None" |
83 |
||
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
84 |
fun lift_un :: "float interval option \<Rightarrow> (float interval \<Rightarrow> float interval option) \<Rightarrow> float interval option" where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
85 |
"lift_un (Some ivl) f = f ivl" | |
29805 | 86 |
"lift_un b f = None" |
87 |
||
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
88 |
lemma lift_un_eq:\<comment> \<open>TODO\<close> "lift_un x f = Option.bind x f" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
89 |
by (cases x) auto |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
90 |
|
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
91 |
fun lift_un' :: "(float interval) option \<Rightarrow> (float interval \<Rightarrow> (float interval)) \<Rightarrow> (float interval) option" where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
92 |
"lift_un' (Some ivl1) f = Some (f ivl1)" | |
29805 | 93 |
"lift_un' b f = None" |
94 |
||
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
95 |
definition bounded_by :: "real list \<Rightarrow> (float interval) option list \<Rightarrow> bool" where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
96 |
"bounded_by xs vs \<longleftrightarrow> (\<forall> i < length vs. case vs ! i of None \<Rightarrow> True | Some ivl \<Rightarrow> xs ! i \<in>\<^sub>r ivl)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
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 |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
101 |
| Some ivl \<Rightarrow> xs ! i \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
102 |
using assms |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
103 |
by (auto simp: bounded_by_def) |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
104 |
|
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
105 |
lemma bounded_by_update: |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
106 |
assumes "bounded_by xs vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
107 |
and bnd: "xs ! i \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
108 |
shows "bounded_by xs (vs[i := Some ivl])" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
109 |
using assms |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
110 |
by (cases "i < length vs") (auto simp: bounded_by_def nth_list_update split: option.splits) |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
111 |
|
60680 | 112 |
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
|
113 |
unfolding bounded_by_def by auto |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
114 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
115 |
fun approx approx' :: "nat \<Rightarrow> floatarith \<Rightarrow> (float interval) option list \<Rightarrow> (float interval) option" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
116 |
where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
117 |
"approx' prec a bs = (case (approx prec a bs) of Some ivl \<Rightarrow> Some (round_interval prec ivl) | None \<Rightarrow> None)" | |
73537
56db8559eadb
fixed problematic addition operation in the 'approximation' package (previous version used much too high precision sometimes)
Manuel Eberl <eberlm@in.tum.de>
parents:
71037
diff
changeset
|
118 |
"approx prec (Add a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (plus_float_interval prec)" | |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
119 |
"approx prec (Minus a) bs = lift_un' (approx' prec a bs) uminus" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
120 |
"approx prec (Mult a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (mult_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
121 |
"approx prec (Inverse a) bs = lift_un (approx' prec a bs) (inverse_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
122 |
"approx prec (Cos a) bs = lift_un' (approx' prec a bs) (cos_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
123 |
"approx prec Pi bs = Some (pi_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
124 |
"approx prec (Min a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (min_interval)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
125 |
"approx prec (Max a b) bs = lift_bin' (approx' prec a bs) (approx' prec b bs) (max_interval)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
126 |
"approx prec (Abs a) bs = lift_un' (approx' prec a bs) (abs_interval)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
127 |
"approx prec (Arctan a) bs = lift_un' (approx' prec a bs) (arctan_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
128 |
"approx prec (Sqrt a) bs = lift_un' (approx' prec a bs) (sqrt_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
129 |
"approx prec (Exp a) bs = lift_un' (approx' prec a bs) (exp_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
130 |
"approx prec (Powr a b) bs = lift_bin (approx' prec a bs) (approx' prec b bs) (powr_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
131 |
"approx prec (Ln a) bs = lift_un (approx' prec a bs) (ln_float_interval prec)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
132 |
"approx prec (Power a n) bs = lift_un' (approx' prec a bs) (power_float_interval prec n)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
133 |
"approx prec (Floor a) bs = lift_un' (approx' prec a bs) (floor_float_interval)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
134 |
"approx prec (Num f) bs = Some (interval_of f)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
135 |
"approx prec (Var i) bs = (if i < length bs then bs ! i else None)" |
29805 | 136 |
|
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
137 |
lemma approx_approx': |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
138 |
assumes Pa: "\<And>ivl. approx prec a vs = Some ivl \<Longrightarrow> interpret_floatarith a xs \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
139 |
and approx': "approx' prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
140 |
shows "interpret_floatarith a xs \<in>\<^sub>r ivl" |
29805 | 141 |
proof - |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
142 |
obtain ivl' where S: "approx prec a vs = Some ivl'" and ivl_def: "ivl = round_interval prec ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
143 |
using approx' unfolding approx'.simps by (cases "approx prec a vs") auto |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
144 |
show ?thesis |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
145 |
by (auto simp: ivl_def intro!: in_round_intervalI Pa[OF S]) |
29805 | 146 |
qed |
147 |
||
148 |
lemma approx: |
|
149 |
assumes "bounded_by xs vs" |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
150 |
and "approx prec arith vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
151 |
shows "interpret_floatarith arith xs \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
152 |
using \<open>approx prec arith vs = Some ivl\<close> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
153 |
using \<open>bounded_by xs vs\<close>[THEN bounded_byE] |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
154 |
by (induct arith arbitrary: ivl) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
155 |
(force split: option.splits if_splits |
73537
56db8559eadb
fixed problematic addition operation in the 'approximation' package (previous version used much too high precision sometimes)
Manuel Eberl <eberlm@in.tum.de>
parents:
71037
diff
changeset
|
156 |
intro!: plus_float_intervalI mult_float_intervalI uminus_in_float_intervalI |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
157 |
inverse_float_interval_eqI abs_in_float_intervalI |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
158 |
min_intervalI max_intervalI |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
159 |
cos_float_intervalI |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
160 |
arctan_float_intervalI pi_float_interval sqrt_float_intervalI exp_float_intervalI |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
161 |
powr_float_interval_eqI ln_float_interval_eqI power_float_intervalI floor_float_intervalI |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
162 |
intro: in_round_intervalI)+ |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
163 |
|
58310 | 164 |
datatype form = Bound floatarith floatarith floatarith form |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
165 |
| Assign floatarith floatarith form |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
166 |
| Less floatarith floatarith |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
167 |
| LessEqual floatarith floatarith |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
168 |
| AtLeastAtMost floatarith floatarith floatarith |
58986 | 169 |
| Conj form form |
170 |
| Disj form form |
|
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
171 |
|
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
172 |
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
|
173 |
"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
|
174 |
"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
|
175 |
"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
|
176 |
"interpret_form (LessEqual a b) vs = (interpret_floatarith a vs \<le> interpret_floatarith b vs)" | |
58986 | 177 |
"interpret_form (AtLeastAtMost x a b) vs = (interpret_floatarith x vs \<in> { interpret_floatarith a vs .. interpret_floatarith b vs })" | |
178 |
"interpret_form (Conj f g) vs \<longleftrightarrow> interpret_form f vs \<and> interpret_form g vs" | |
|
179 |
"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
|
180 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
181 |
fun approx_form' and approx_form :: "nat \<Rightarrow> form \<Rightarrow> (float interval) option list \<Rightarrow> nat list \<Rightarrow> bool" where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
182 |
"approx_form' prec f 0 n ivl bs ss = approx_form prec f (bs[n := Some ivl]) ss" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
183 |
"approx_form' prec f (Suc s) n ivl bs ss = |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
184 |
(let (ivl1, ivl2) = split_float_interval ivl |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
185 |
in (if approx_form' prec f s n ivl1 bs ss then approx_form' prec f s n ivl2 bs ss else 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
|
186 |
"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
|
187 |
(case (approx prec a bs, approx prec b bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
188 |
of (Some ivl1, Some ivl2) \<Rightarrow> approx_form' prec f (ss ! n) n (sup ivl1 ivl2) bs ss |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
189 |
| _ \<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
|
190 |
"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
|
191 |
(case (approx prec a bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
192 |
of (Some ivl) \<Rightarrow> approx_form' prec f (ss ! n) n ivl bs ss |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
193 |
| _ \<Rightarrow> False)" | |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
194 |
"approx_form prec (Less a b) bs ss = |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
195 |
(case (approx prec a bs, approx prec b bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
196 |
of (Some ivl, Some ivl') \<Rightarrow> float_plus_up prec (upper ivl) (-lower ivl') < 0 |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
197 |
| _ \<Rightarrow> False)" | |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
198 |
"approx_form prec (LessEqual a b) bs ss = |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
199 |
(case (approx prec a bs, approx prec b bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
200 |
of (Some ivl, Some ivl') \<Rightarrow> float_plus_up prec (upper ivl) (-lower ivl') \<le> 0 |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
201 |
| _ \<Rightarrow> False)" | |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
202 |
"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
|
203 |
(case (approx prec x bs, approx prec a bs, approx prec b bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
204 |
of (Some ivlx, Some ivl, Some ivl') \<Rightarrow> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
205 |
float_plus_up prec (upper ivl) (-lower ivlx) \<le> 0 \<and> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
206 |
float_plus_up prec (upper ivlx) (-lower ivl') \<le> 0 |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
207 |
| _ \<Rightarrow> False)" | |
58986 | 208 |
"approx_form prec (Conj a b) bs ss \<longleftrightarrow> approx_form prec a bs ss \<and> approx_form prec b bs ss" | |
209 |
"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
|
210 |
"approx_form _ _ _ _ = False" |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
211 |
|
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
|
212 |
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
|
213 |
|
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
214 |
lemma approx_form_approx_form': |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
215 |
assumes "approx_form' prec f s n ivl bs ss" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
216 |
and "(x::real) \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
217 |
obtains ivl' where "x \<in>\<^sub>r ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
218 |
and "approx_form prec f (bs[n := Some ivl']) ss" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
219 |
using assms proof (induct s arbitrary: ivl) |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
220 |
case 0 |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
221 |
from this(1)[of ivl] this(2,3) |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
222 |
show thesis by auto |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
223 |
next |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
224 |
case (Suc s) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
225 |
then obtain ivl1 ivl2 where ivl_split: "split_float_interval ivl = (ivl1, ivl2)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
226 |
by (fastforce dest: split_float_intervalD) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
227 |
from split_float_interval_realD[OF this \<open>x \<in>\<^sub>r ivl\<close>] |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
228 |
consider "x \<in>\<^sub>r ivl1" | "x \<in>\<^sub>r ivl2" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
229 |
by atomize_elim |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
230 |
then show thesis |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
231 |
proof cases |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
232 |
case *: 1 |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
233 |
from Suc.hyps[OF _ _ *] Suc.prems |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
234 |
show ?thesis |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
235 |
by (simp add: lazy_conj ivl_split split: prod.splits) |
29805 | 236 |
next |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
237 |
case *: 2 |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
238 |
from Suc.hyps[OF _ _ *] Suc.prems |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
239 |
show ?thesis |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
240 |
by (simp add: lazy_conj ivl_split split: prod.splits) |
29805 | 241 |
qed |
242 |
qed |
|
243 |
||
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
244 |
lemma approx_form_aux: |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
245 |
assumes "approx_form prec f vs ss" |
49351 | 246 |
and "bounded_by xs vs" |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
247 |
shows "interpret_form f xs" |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
248 |
using assms proof (induct f arbitrary: vs) |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
249 |
case (Bound x a b f) |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
250 |
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
|
251 |
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
|
252 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
253 |
with Bound.prems obtain ivl1 ivl2 |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
254 |
where l_eq: "approx prec a vs = Some ivl1" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
255 |
and u_eq: "approx prec b vs = Some ivl2" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
256 |
and approx_form': "approx_form' prec f (ss ! n) n (sup ivl1 ivl2) vs ss" |
37411
c88c44156083
removed simplifier congruence rule of "prod_case"
haftmann
parents:
37391
diff
changeset
|
257 |
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
|
258 |
|
60680 | 259 |
have "interpret_form f xs" |
260 |
if "xs ! n \<in> { interpret_floatarith a xs .. interpret_floatarith b xs }" |
|
261 |
proof - |
|
262 |
from approx[OF Bound.prems(2) l_eq] and approx[OF Bound.prems(2) u_eq] that |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
263 |
have "xs ! n \<in>\<^sub>r (sup ivl1 ivl2)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
264 |
by (auto simp: set_of_eq sup_float_def max_def inf_float_def min_def) |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
265 |
|
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
266 |
from approx_form_approx_form'[OF approx_form' this] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
267 |
obtain ivlx where bnds: "xs ! n \<in>\<^sub>r ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
268 |
and approx_form: "approx_form prec f (vs[n := Some ivlx]) ss" . |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
269 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
270 |
from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some ivlx])" |
60680 | 271 |
by (rule bounded_by_update) |
272 |
with Bound.hyps[OF approx_form] show ?thesis |
|
273 |
by blast |
|
274 |
qed |
|
275 |
thus ?case |
|
276 |
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
|
277 |
next |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
278 |
case (Assign x a f) |
60680 | 279 |
then obtain n where x_eq: "x = Var n" |
280 |
by (cases x) auto |
|
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
281 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
282 |
with Assign.prems obtain ivl |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
283 |
where bnd_eq: "approx prec a vs = Some ivl" |
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
|
284 |
and x_eq: "x = Var n" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
285 |
and approx_form': "approx_form' prec f (ss ! n) n ivl vs ss" |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
286 |
by (cases "approx prec a vs") auto |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
287 |
|
60680 | 288 |
have "interpret_form f xs" |
289 |
if bnds: "xs ! n = interpret_floatarith a xs" |
|
290 |
proof - |
|
291 |
from approx[OF Assign.prems(2) bnd_eq] bnds |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
292 |
have "xs ! n \<in>\<^sub>r ivl" by auto |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
293 |
from approx_form_approx_form'[OF approx_form' this] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
294 |
obtain ivlx where bnds: "xs ! n \<in>\<^sub>r ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
295 |
and approx_form: "approx_form prec f (vs[n := Some ivlx]) ss" . |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
296 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
297 |
from \<open>bounded_by xs vs\<close> bnds have "bounded_by xs (vs[n := Some ivlx])" |
60680 | 298 |
by (rule bounded_by_update) |
299 |
with Assign.hyps[OF approx_form] show ?thesis |
|
300 |
by blast |
|
301 |
qed |
|
302 |
thus ?case |
|
303 |
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
|
304 |
next |
29805 | 305 |
case (Less a b) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
306 |
then obtain ivl ivl' |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
307 |
where l_eq: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
308 |
and u_eq: "approx prec b vs = Some ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
309 |
and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivl')) < 0" |
60680 | 310 |
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
|
311 |
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
|
312 |
approx[OF Less.prems(2) l_eq] approx[OF Less.prems(2) u_eq] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
313 |
show ?case by (auto simp: set_of_eq) |
29805 | 314 |
next |
315 |
case (LessEqual a b) |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
316 |
then obtain ivl ivl' |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
317 |
where l_eq: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
318 |
and u_eq: "approx prec b vs = Some ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
319 |
and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivl')) \<le> 0" |
60680 | 320 |
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
|
321 |
from order_trans[OF float_plus_up inequality] |
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
immler
parents:
58982
diff
changeset
|
322 |
approx[OF LessEqual.prems(2) l_eq] approx[OF LessEqual.prems(2) u_eq] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
323 |
show ?case by (auto simp: set_of_eq) |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
324 |
next |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
325 |
case (AtLeastAtMost x a b) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
326 |
then obtain ivlx ivl ivl' |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
327 |
where x_eq: "approx prec x vs = Some ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
328 |
and l_eq: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
329 |
and u_eq: "approx prec b vs = Some ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
330 |
and inequality: "real_of_float (float_plus_up prec (upper ivl) (-lower ivlx)) \<le> 0" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
331 |
"real_of_float (float_plus_up prec (upper ivlx) (-lower ivl')) \<le> 0" |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
332 |
by (cases "approx prec x vs", auto, |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
333 |
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
|
334 |
cases "approx prec b vs", auto) |
58985
bf498e0af9e3
truncate intermediate results in horner to improve performance of approximate;
immler
parents:
58982
diff
changeset
|
335 |
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
|
336 |
approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
337 |
show ?case by (auto simp: set_of_eq) |
58986 | 338 |
qed auto |
29805 | 339 |
|
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
340 |
lemma approx_form: |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
341 |
assumes "n = length xs" |
60680 | 342 |
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
|
343 |
shows "interpret_form f xs" |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
344 |
using approx_form_aux[OF _ bounded_by_None] assms by auto |
29805 | 345 |
|
60680 | 346 |
|
60533 | 347 |
subsection \<open>Implementing Taylor series expansion\<close> |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
348 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
349 |
fun isDERIV :: "nat \<Rightarrow> floatarith \<Rightarrow> real list \<Rightarrow> bool" where |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
350 |
"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
|
351 |
"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
|
352 |
"isDERIV x (Minus a) vs = isDERIV x a vs" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
353 |
"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
|
354 |
"isDERIV x (Cos a) vs = isDERIV x a vs" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
355 |
"isDERIV x (Arctan a) vs = isDERIV x a vs" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
356 |
"isDERIV x (Min a b) vs = False" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
357 |
"isDERIV x (Max a b) vs = False" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
358 |
"isDERIV x (Abs a) vs = False" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
359 |
"isDERIV x Pi vs = True" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
360 |
"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
|
361 |
"isDERIV x (Exp a) vs = isDERIV x a vs" | |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
362 |
"isDERIV x (Powr a b) vs = |
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
363 |
(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
|
364 |
"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
|
365 |
"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
|
366 |
"isDERIV x (Power a 0) vs = True" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
367 |
"isDERIV x (Power a (Suc n)) vs = isDERIV x a vs" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
368 |
"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
|
369 |
"isDERIV x (Var n) vs = True" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
370 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
371 |
fun DERIV_floatarith :: "nat \<Rightarrow> floatarith \<Rightarrow> floatarith" where |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
372 |
"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
|
373 |
"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
|
374 |
"DERIV_floatarith x (Minus a) = Minus (DERIV_floatarith x a)" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
375 |
"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
|
376 |
"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
|
377 |
"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
|
378 |
"DERIV_floatarith x (Min a b) = Num 0" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
379 |
"DERIV_floatarith x (Max a b) = Num 0" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
380 |
"DERIV_floatarith x (Abs a) = Num 0" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
381 |
"DERIV_floatarith x Pi = Num 0" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
382 |
"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
|
383 |
"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
|
384 |
"DERIV_floatarith x (Powr a b) = |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
385 |
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
|
386 |
"DERIV_floatarith x (Ln a) = Mult (Inverse a) (DERIV_floatarith x a)" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
387 |
"DERIV_floatarith x (Power a 0) = Num 0" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
388 |
"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
|
389 |
"DERIV_floatarith x (Floor a) = Num 0" | |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
390 |
"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
|
391 |
"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
|
392 |
|
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
393 |
lemma has_real_derivative_powr': |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
394 |
fixes f g :: "real \<Rightarrow> real" |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
395 |
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
|
396 |
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
|
397 |
assumes "f x > 0" |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
398 |
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
|
399 |
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
|
400 |
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
|
401 |
from assms have "isCont f x" |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
402 |
by (simp add: DERIV_continuous) |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
403 |
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
|
404 |
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
|
405 |
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
|
406 |
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
|
407 |
by eventually_elim (simp add: powr_def) |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
408 |
next |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
409 |
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
|
410 |
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
|
411 |
qed |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
412 |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
413 |
lemma DERIV_floatarith: |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
414 |
assumes "n < length vs" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
415 |
assumes isDERIV: "isDERIV n f (vs[n := x])" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
416 |
shows "DERIV (\<lambda> x'. interpret_floatarith f (vs[n := x'])) x :> |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
417 |
interpret_floatarith (DERIV_floatarith n f) (vs[n := x])" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
418 |
(is "DERIV (?i f) x :> _") |
49351 | 419 |
using isDERIV |
420 |
proof (induct f arbitrary: x) |
|
421 |
case (Inverse a) |
|
422 |
thus ?case |
|
56381
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents:
56195
diff
changeset
|
423 |
by (auto intro!: derivative_eq_intros simp add: algebra_simps power2_eq_square) |
49351 | 424 |
next |
425 |
case (Cos a) |
|
426 |
thus ?case |
|
56382 | 427 |
by (auto intro!: derivative_eq_intros |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
428 |
simp del: interpret_floatarith.simps(5) |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
429 |
simp add: interpret_floatarith_sin interpret_floatarith.simps(5)[of a]) |
49351 | 430 |
next |
431 |
case (Power a n) |
|
432 |
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
|
433 |
by (cases n) (auto intro!: derivative_eq_intros simp del: power_Suc) |
49351 | 434 |
next |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
435 |
case (Floor a) |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
436 |
thus ?case |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
437 |
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
|
438 |
next |
49351 | 439 |
case (Ln a) |
56382 | 440 |
thus ?case by (auto intro!: derivative_eq_intros simp add: divide_inverse) |
49351 | 441 |
next |
442 |
case (Var i) |
|
60533 | 443 |
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
|
444 |
next |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
445 |
case (Powr a b) |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
446 |
note [derivative_intros] = has_real_derivative_powr' |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
447 |
from Powr show ?case |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
448 |
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
|
449 |
qed (auto intro!: derivative_eq_intros) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
450 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
451 |
declare approx.simps[simp del] |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
452 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
453 |
fun isDERIV_approx :: "nat \<Rightarrow> nat \<Rightarrow> floatarith \<Rightarrow> (float interval) option list \<Rightarrow> bool" where |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
454 |
"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
|
455 |
"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
|
456 |
"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
|
457 |
"isDERIV_approx prec x (Inverse a) vs = |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
458 |
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl \<or> upper ivl < 0 | None \<Rightarrow> False))" | |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
459 |
"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
|
460 |
"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
|
461 |
"isDERIV_approx prec x (Min a b) vs = False" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
462 |
"isDERIV_approx prec x (Max a b) vs = False" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
463 |
"isDERIV_approx prec x (Abs a) vs = False" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
464 |
"isDERIV_approx prec x Pi vs = True" | |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
465 |
"isDERIV_approx prec x (Sqrt a) vs = |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
466 |
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl | None \<Rightarrow> False))" | |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
467 |
"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
|
468 |
"isDERIV_approx prec x (Powr a b) vs = |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
469 |
(isDERIV_approx prec x a vs \<and> isDERIV_approx prec x b vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl | None \<Rightarrow> False))" | |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
470 |
"isDERIV_approx prec x (Ln a) vs = |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
471 |
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> 0 < lower ivl | None \<Rightarrow> False))" | |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
472 |
"isDERIV_approx prec x (Power a 0) vs = True" | |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
473 |
"isDERIV_approx prec x (Floor a) vs = |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
474 |
(isDERIV_approx prec x a vs \<and> (case approx prec a vs of Some ivl \<Rightarrow> lower ivl > floor (upper ivl) \<and> upper ivl < ceiling (lower ivl) | None \<Rightarrow> False))" | |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
475 |
"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
|
476 |
"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
|
477 |
"isDERIV_approx prec x (Var n) vs = True" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
478 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
479 |
lemma isDERIV_approx: |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
480 |
assumes "bounded_by xs vs" |
49351 | 481 |
and isDERIV_approx: "isDERIV_approx prec x f vs" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
482 |
shows "isDERIV x f xs" |
49351 | 483 |
using isDERIV_approx |
484 |
proof (induct f) |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
485 |
case (Inverse a) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
486 |
then obtain ivl where approx_Some: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
487 |
and *: "0 < lower ivl \<or> upper ivl < 0" |
49351 | 488 |
by (cases "approx prec a vs") auto |
60533 | 489 |
with approx[OF \<open>bounded_by xs vs\<close> approx_Some] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
490 |
have "interpret_floatarith a xs \<noteq> 0" by (auto simp: set_of_eq) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
491 |
thus ?case using Inverse by auto |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
492 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
493 |
case (Ln a) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
494 |
then obtain ivl where approx_Some: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
495 |
and *: "0 < lower ivl" |
49351 | 496 |
by (cases "approx prec a vs") auto |
60533 | 497 |
with approx[OF \<open>bounded_by xs vs\<close> approx_Some] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
498 |
have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
499 |
thus ?case using Ln by auto |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
500 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
501 |
case (Sqrt a) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
502 |
then obtain ivl where approx_Some: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
503 |
and *: "0 < lower ivl" |
49351 | 504 |
by (cases "approx prec a vs") auto |
60533 | 505 |
with approx[OF \<open>bounded_by xs vs\<close> approx_Some] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
506 |
have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
507 |
thus ?case using Sqrt by auto |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
508 |
next |
60680 | 509 |
case (Power a n) |
510 |
thus ?case by (cases n) auto |
|
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
511 |
next |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
512 |
case (Powr a b) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
513 |
from Powr obtain ivl1 where a: "approx prec a vs = Some ivl1" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
514 |
and pos: "0 < lower ivl1" |
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
515 |
by (cases "approx prec a vs") auto |
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
516 |
with approx[OF \<open>bounded_by xs vs\<close> a] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
517 |
have "0 < interpret_floatarith a xs" by (auto simp: set_of_eq) |
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
518 |
with Powr show ?case by auto |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
519 |
next |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
520 |
case (Floor a) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
521 |
then obtain ivl where approx_Some: "approx prec a vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
522 |
and "real_of_int \<lfloor>real_of_float (upper ivl)\<rfloor> < real_of_float (lower ivl)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
523 |
"real_of_float (upper ivl) < real_of_int \<lceil>real_of_float (lower ivl)\<rceil>" |
63263
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
524 |
and "isDERIV x a xs" |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
525 |
by (cases "approx prec a vs") auto |
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
immler
parents:
63248
diff
changeset
|
526 |
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
|
527 |
show ?case |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
528 |
by (force elim!: Ints_cases simp: set_of_eq) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
529 |
qed auto |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
530 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
531 |
lemma bounded_by_update_var: |
60680 | 532 |
assumes "bounded_by xs vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
533 |
and "vs ! i = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
534 |
and bnd: "x \<in>\<^sub>r ivl" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
535 |
shows "bounded_by (xs[i := x]) vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
536 |
using assms |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
537 |
using nth_list_update |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
538 |
by (cases "i < length xs") |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
539 |
(force simp: bounded_by_def split: option.splits)+ |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
540 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
541 |
lemma isDERIV_approx': |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
542 |
assumes "bounded_by xs vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
543 |
and vs_x: "vs ! x = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
544 |
and X_in: "X \<in>\<^sub>r ivl" |
49351 | 545 |
and approx: "isDERIV_approx prec x f vs" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
546 |
shows "isDERIV x f (xs[x := X])" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
547 |
proof - |
60680 | 548 |
from bounded_by_update_var[OF \<open>bounded_by xs vs\<close> vs_x X_in] approx |
549 |
show ?thesis by (rule isDERIV_approx) |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
550 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
551 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
552 |
lemma DERIV_approx: |
60680 | 553 |
assumes "n < length xs" |
554 |
and bnd: "bounded_by xs vs" |
|
49351 | 555 |
and isD: "isDERIV_approx prec n f vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
556 |
and app: "approx prec (DERIV_floatarith n f) vs = Some ivl" (is "approx _ ?D _ = _") |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
557 |
shows "\<exists>(x::real). x \<in>\<^sub>r ivl \<and> |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
558 |
DERIV (\<lambda> x. interpret_floatarith f (xs[n := x])) (xs!n) :> x" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
559 |
(is "\<exists> x. _ \<and> DERIV (?i f) _ :> _") |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
560 |
proof (rule exI[of _ "?i ?D (xs!n)"], rule conjI) |
60680 | 561 |
let "?i f" = "\<lambda>x. interpret_floatarith f (xs[n := x])" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
562 |
from approx[OF bnd app] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
563 |
show "?i ?D (xs!n) \<in>\<^sub>r ivl" |
60533 | 564 |
using \<open>n < length xs\<close> by auto |
565 |
from DERIV_floatarith[OF \<open>n < length xs\<close>, of f "xs!n"] isDERIV_approx[OF bnd isD] |
|
60680 | 566 |
show "DERIV (?i f) (xs!n) :> (?i ?D (xs!n))" |
567 |
by simp |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
568 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
569 |
|
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
570 |
lemma lift_bin_aux: |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
571 |
assumes lift_bin_Some: "lift_bin a b f = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
572 |
obtains ivl1 ivl2 |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
573 |
where "a = Some ivl1" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
574 |
and "b = Some ivl2" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
575 |
and "f ivl1 ivl2 = Some ivl" |
49351 | 576 |
using assms by (cases a, simp, cases b, simp, auto) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
577 |
|
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
578 |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
579 |
fun approx_tse where |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
580 |
"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
|
581 |
"approx_tse prec n (Suc s) c k f bs = |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
582 |
(if isDERIV_approx prec n f bs then |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
583 |
lift_bin (approx prec f (bs[n := Some (interval_of c)])) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
584 |
(approx_tse prec n s c (Suc k) (DERIV_floatarith n f) bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
585 |
(\<lambda>ivl1 ivl2. 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
|
586 |
(Add (Var 0) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
587 |
(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
|
588 |
(Mult (Add (Var (Suc (Suc 0))) (Minus (Num c))) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
589 |
(Var (Suc 0))))) [Some ivl1, Some ivl2, bs!n]) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
590 |
else approx prec f bs)" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
591 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
592 |
lemma bounded_by_Cons: |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
593 |
assumes bnd: "bounded_by xs vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
594 |
and x: "x \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
595 |
shows "bounded_by (x#xs) ((Some ivl)#vs)" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
596 |
proof - |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
597 |
have "case ((Some ivl)#vs) ! i of Some ivl \<Rightarrow> (x#xs)!i \<in>\<^sub>r ivl | None \<Rightarrow> True" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
598 |
if *: "i < length ((Some ivl)#vs)" for i |
60680 | 599 |
proof (cases i) |
600 |
case 0 |
|
601 |
with x show ?thesis by auto |
|
602 |
next |
|
603 |
case (Suc i) |
|
604 |
with * have "i < length vs" by auto |
|
605 |
from bnd[THEN bounded_byE, OF this] |
|
606 |
show ?thesis unfolding Suc nth_Cons_Suc . |
|
607 |
qed |
|
608 |
thus ?thesis |
|
609 |
by (auto simp add: bounded_by_def) |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
610 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
611 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
612 |
lemma approx_tse_generic: |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
613 |
assumes "bounded_by xs vs" |
60680 | 614 |
and bnd_c: "bounded_by (xs[x := c]) vs" |
615 |
and "x < length vs" and "x < length xs" |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
616 |
and bnd_x: "vs ! x = Some ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
617 |
and ate: "approx_tse prec x s c k f vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
618 |
shows "\<exists> n. (\<forall> m < n. \<forall> (z::real) \<in> set_of (real_interval ivlx). |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
619 |
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
|
620 |
(interpret_floatarith ((DERIV_floatarith x ^^ (Suc m)) f) (xs[x := z]))) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
621 |
\<and> (\<forall> (t::real) \<in> set_of (real_interval ivlx). (\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) * |
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
622 |
interpret_floatarith ((DERIV_floatarith x ^^ i) f) (xs[x := c]) * |
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
623 |
(xs!x - c)^i) + |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
624 |
inverse (real (\<Prod> j \<in> {k..<k+n}. j)) * |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
625 |
interpret_floatarith ((DERIV_floatarith x ^^ n) f) (xs[x := t]) * |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
626 |
(xs!x - c)^n \<in>\<^sub>r ivl)" (is "\<exists> n. ?taylor f k ivl n") |
60680 | 627 |
using ate |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
628 |
proof (induct s arbitrary: k f ivl) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
629 |
case 0 |
49351 | 630 |
{ |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
631 |
fix t::real assume "t \<in>\<^sub>r ivlx" |
60533 | 632 |
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
|
633 |
from approx[OF this 0[unfolded approx_tse.simps]] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
634 |
have "(interpret_floatarith f (xs[x := t])) \<in>\<^sub>r ivl" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
635 |
by (auto simp add: algebra_simps) |
49351 | 636 |
} |
637 |
thus ?case by (auto intro!: exI[of _ 0]) |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
638 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
639 |
case (Suc s) |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
640 |
show ?case |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
641 |
proof (cases "isDERIV_approx prec x f vs") |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
642 |
case False |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
643 |
note ap = Suc.prems[unfolded approx_tse.simps if_not_P[OF False]] |
49351 | 644 |
{ |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
645 |
fix t::real assume "t \<in>\<^sub>r ivlx" |
60533 | 646 |
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
|
647 |
from approx[OF this ap] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
648 |
have "(interpret_floatarith f (xs[x := t])) \<in>\<^sub>r ivl" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32920
diff
changeset
|
649 |
by (auto simp add: algebra_simps) |
49351 | 650 |
} |
651 |
thus ?thesis by (auto intro!: exI[of _ 0]) |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
652 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
653 |
case True |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
654 |
with Suc.prems |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
655 |
obtain ivl1 ivl2 |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
656 |
where a: "approx prec f (vs[x := Some (interval_of c)]) = Some ivl1" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
657 |
and ate: "approx_tse prec x s c (Suc k) (DERIV_floatarith x f) vs = Some ivl2" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
658 |
and final: "approx prec |
49351 | 659 |
(Add (Var 0) |
660 |
(Mult (Inverse (Num (Float (int k) 0))) |
|
661 |
(Mult (Add (Var (Suc (Suc 0))) (Minus (Num c))) |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
662 |
(Var (Suc 0))))) [Some ivl1, Some ivl2, vs!x] = Some ivl" |
62200
67792e4a5486
Made Approximation work for powr again
Manuel Eberl <eberlm@in.tum.de>
parents:
61969
diff
changeset
|
663 |
by (auto elim!: lift_bin_aux) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
664 |
|
60533 | 665 |
from bnd_c \<open>x < length xs\<close> |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
666 |
have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (interval_of c)])" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
667 |
by (auto intro!: bounded_by_update) |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
668 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
669 |
from approx[OF this a] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
670 |
have f_c: "interpret_floatarith ((DERIV_floatarith x ^^ 0) f) (xs[x := c]) \<in>\<^sub>r ivl1" |
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
|
671 |
(is "?f 0 (real_of_float c) \<in> _") |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
672 |
by auto |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
673 |
|
60680 | 674 |
have funpow_Suc[symmetric]: "(f ^^ Suc n) x = (f ^^ n) (f x)" |
675 |
for f :: "'a \<Rightarrow> 'a" and n :: nat and x :: 'a |
|
676 |
by (induct n) auto |
|
677 |
from Suc.hyps[OF ate, unfolded this] obtain n |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
678 |
where DERIV_hyp: "\<And>m z. \<lbrakk> m < n ; (z::real) \<in>\<^sub>r ivlx \<rbrakk> \<Longrightarrow> |
60680 | 679 |
DERIV (?f (Suc m)) z :> ?f (Suc (Suc m)) z" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
680 |
and hyp: "\<forall>t \<in> set_of (real_interval ivlx). |
60680 | 681 |
(\<Sum> i = 0..<n. inverse (real (\<Prod> j \<in> {Suc k..<Suc k + i}. j)) * ?f (Suc i) c * (xs!x - c)^i) + |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
682 |
inverse (real (\<Prod> j \<in> {Suc k..<Suc k + n}. j)) * ?f (Suc n) t * (xs!x - c)^n \<in>\<^sub>r ivl2" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
683 |
(is "\<forall> t \<in> _. ?X (Suc k) f n t \<in> _") |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
684 |
by blast |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
685 |
|
60680 | 686 |
have DERIV: "DERIV (?f m) z :> ?f (Suc m) z" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
687 |
if "m < Suc n" and bnd_z: "z \<in>\<^sub>r ivlx" for m and z::real |
60680 | 688 |
proof (cases m) |
689 |
case 0 |
|
690 |
with DERIV_floatarith[OF \<open>x < length xs\<close> |
|
691 |
isDERIV_approx'[OF \<open>bounded_by xs vs\<close> bnd_x bnd_z True]] |
|
692 |
show ?thesis by simp |
|
693 |
next |
|
694 |
case (Suc m') |
|
695 |
hence "m' < n" |
|
696 |
using \<open>m < Suc n\<close> by auto |
|
697 |
from DERIV_hyp[OF this bnd_z] show ?thesis |
|
698 |
using Suc by simp |
|
699 |
qed |
|
700 |
||
701 |
have "\<And>k i. k < i \<Longrightarrow> {k ..< i} = insert k {Suc k ..< i}" by auto |
|
64272 | 702 |
hence prod_head_Suc: "\<And>k i. \<Prod>{k ..< k + Suc i} = k * \<Prod>{Suc k ..< Suc k + i}" |
60680 | 703 |
by auto |
64267 | 704 |
have sum_move0: "\<And>k F. sum F {0..<Suc k} = F 0 + sum (\<lambda> k. F (Suc k)) {0..<k}" |
70113
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents:
70097
diff
changeset
|
705 |
unfolding sum.shift_bounds_Suc_ivl[symmetric] |
70097
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
69835
diff
changeset
|
706 |
unfolding sum.atLeast_Suc_lessThan[OF zero_less_Suc] .. |
63040 | 707 |
define C where "C = xs!x - c" |
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
708 |
|
49351 | 709 |
{ |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
710 |
fix t::real assume t: "t \<in>\<^sub>r ivlx" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
711 |
hence "bounded_by [xs!x] [vs!x]" |
60533 | 712 |
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
|
713 |
by (cases "vs!x", auto simp add: bounded_by_def) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
714 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
715 |
with hyp[THEN bspec, OF t] f_c |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
716 |
have "bounded_by [?f 0 c, ?X (Suc k) f n t, xs!x] [Some ivl1, Some ivl2, vs!x]" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32920
diff
changeset
|
717 |
by (auto intro!: bounded_by_Cons) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
718 |
from approx[OF this final, unfolded atLeastAtMost_iff[symmetric]] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
719 |
have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse k + ?f 0 c \<in>\<^sub>r ivl" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32920
diff
changeset
|
720 |
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
|
721 |
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
|
722 |
(\<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
|
723 |
inverse (real (\<Prod> j \<in> {k..<k+Suc n}. j)) * ?f (Suc n) t * (xs!x - c)^Suc n" (is "_ = ?T") |
64272 | 724 |
unfolding funpow_Suc C_def[symmetric] sum_move0 prod_head_Suc |
35082 | 725 |
by (auto simp add: algebra_simps) |
64267 | 726 |
(simp only: mult.left_commute [of _ "inverse (real k)"] sum_distrib_left [symmetric]) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
727 |
finally have "?T \<in>\<^sub>r ivl" . |
49351 | 728 |
} |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
729 |
thus ?thesis using DERIV by blast |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
730 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
731 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
732 |
|
64272 | 733 |
lemma prod_fact: "real (\<Prod> {1..<1 + k}) = fact (k :: nat)" |
734 |
by (simp add: fact_prod atLeastLessThanSuc_atLeastAtMost) |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
735 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
736 |
lemma approx_tse: |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
737 |
assumes "bounded_by xs vs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
738 |
and bnd_x: "vs ! x = Some ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
739 |
and bnd_c: "real_of_float c \<in>\<^sub>r ivlx" |
49351 | 740 |
and "x < length vs" and "x < length xs" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
741 |
and ate: "approx_tse prec x s c 1 f vs = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
742 |
shows "interpret_floatarith f xs \<in>\<^sub>r ivl" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
743 |
proof - |
63040 | 744 |
define F where [abs_def]: "F n z = |
745 |
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
|
746 |
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
|
747 |
|
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
748 |
hence "bounded_by (xs[x := c]) vs" and "x < length vs" "x < length xs" |
60533 | 749 |
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
|
750 |
by (auto intro!: bounded_by_update_var) |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
751 |
|
60533 | 752 |
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
|
753 |
obtain n |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
754 |
where DERIV: "\<forall> m z. m < n \<and> real_of_float (lower ivlx) \<le> z \<and> z \<le> real_of_float (upper ivlx) \<longrightarrow> DERIV (F m) z :> F (Suc m) z" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
755 |
and hyp: "\<And> (t::real). t \<in>\<^sub>r ivlx \<Longrightarrow> |
59730
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents:
59658
diff
changeset
|
756 |
(\<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
|
757 |
inverse ((fact n)) * F n t * (xs!x - c)^n |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
758 |
\<in>\<^sub>r ivl" (is "\<And> t. _ \<Longrightarrow> ?taylor t \<in> _") |
64272 | 759 |
unfolding F_def atLeastAtMost_iff[symmetric] prod_fact |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
760 |
by (auto simp: set_of_eq Ball_def) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
761 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
762 |
have bnd_xs: "xs ! x \<in>\<^sub>r ivlx" |
60533 | 763 |
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
|
764 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
765 |
show ?thesis |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
766 |
proof (cases n) |
60680 | 767 |
case 0 |
768 |
thus ?thesis |
|
769 |
using hyp[OF bnd_xs] unfolding F_def by auto |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
770 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
771 |
case (Suc n') |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
772 |
show ?thesis |
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
773 |
proof (cases "xs ! x = c") |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
774 |
case True |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
775 |
from True[symmetric] hyp[OF bnd_xs] Suc show ?thesis |
70113
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents:
70097
diff
changeset
|
776 |
unfolding F_def Suc sum.atLeast_Suc_lessThan[OF zero_less_Suc] sum.shift_bounds_Suc_ivl |
60680 | 777 |
by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
778 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
779 |
case False |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
780 |
have "lower ivlx \<le> real_of_float c" "real_of_float c \<le> upper ivlx" "lower ivlx \<le> xs!x" "xs!x \<le> upper ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
781 |
using Suc bnd_c \<open>bounded_by xs vs\<close>[THEN bounded_byE, OF \<open>x < length vs\<close>] bnd_x |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
782 |
by (auto simp: set_of_eq) |
69529 | 783 |
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
|
784 |
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
|
785 |
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
|
786 |
(\<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
|
787 |
F (Suc n') t / (fact (Suc n')) * (xs ! x - c) ^ Suc n'" |
56195 | 788 |
unfolding atLeast0LessThan by blast |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
789 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
790 |
from t_bnd bnd_xs bnd_c have *: "t \<in>\<^sub>r ivlx" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
791 |
by (cases "xs ! x < c") (auto simp: set_of_eq) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
792 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
793 |
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
|
794 |
unfolding fl_eq Suc by (auto simp add: algebra_simps divide_inverse) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
795 |
also have "\<dots> \<in>\<^sub>r ivl" |
60680 | 796 |
using * by (rule hyp) |
797 |
finally show ?thesis |
|
798 |
by simp |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
799 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
800 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
801 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
802 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
803 |
fun approx_tse_form' where |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
804 |
"approx_tse_form' prec t f 0 ivl cmp = |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
805 |
(case approx_tse prec 0 t (mid ivl) 1 f [Some ivl] |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
806 |
of Some ivl \<Rightarrow> cmp ivl | None \<Rightarrow> False)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
807 |
"approx_tse_form' prec t f (Suc s) ivl cmp = |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
808 |
(let (ivl1, ivl2) = split_float_interval ivl |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
809 |
in (if approx_tse_form' prec t f s ivl1 cmp then |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
810 |
approx_tse_form' prec t f s ivl2 cmp else False))" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
811 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
812 |
lemma approx_tse_form': |
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
813 |
fixes x :: real |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
814 |
assumes "approx_tse_form' prec t f s ivl cmp" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
815 |
and "x \<in>\<^sub>r ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
816 |
shows "\<exists>ivl' ivly. x \<in>\<^sub>r ivl' \<and> ivl' \<le> ivl \<and> cmp ivly \<and> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
817 |
approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly" |
60680 | 818 |
using assms |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
819 |
proof (induct s arbitrary: ivl) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
820 |
case 0 |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
821 |
then obtain ivly |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
822 |
where *: "approx_tse prec 0 t (mid ivl) 1 f [Some ivl] = Some ivly" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
823 |
and **: "cmp ivly" by (auto elim!: case_optionE) |
46545 | 824 |
with 0 show ?case by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
825 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
826 |
case (Suc s) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
827 |
let ?ivl1 = "fst (split_float_interval ivl)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
828 |
let ?ivl2 = "snd (split_float_interval ivl)" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
829 |
from Suc.prems |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
830 |
have l: "approx_tse_form' prec t f s ?ivl1 cmp" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
831 |
and u: "approx_tse_form' prec t f s ?ivl2 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
|
832 |
by (auto simp add: Let_def lazy_conj) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
833 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
834 |
have subintervals: "?ivl1 \<le> ivl" "?ivl2 \<le> ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
835 |
using mid_le |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
836 |
by (auto simp: less_eq_interval_def split_float_interval_bounds) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
837 |
|
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
838 |
from split_float_interval_realD[OF _ \<open>x \<in>\<^sub>r ivl\<close>] consider "x \<in>\<^sub>r ?ivl1" | "x \<in>\<^sub>r ?ivl2" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
839 |
by (auto simp: prod_eq_iff) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
840 |
then show ?case |
60680 | 841 |
proof cases |
842 |
case 1 |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
843 |
from Suc.hyps[OF l this] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
844 |
obtain ivl' ivly where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
845 |
"x \<in>\<^sub>r ivl' \<and> ivl' \<le> fst (split_float_interval ivl) \<and> cmp ivly \<and> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
846 |
approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly" |
60680 | 847 |
by blast |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
848 |
then show ?thesis |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
849 |
using subintervals |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
850 |
by (auto) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
851 |
next |
60680 | 852 |
case 2 |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
853 |
from Suc.hyps[OF u this] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
854 |
obtain ivl' ivly where |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
855 |
"x \<in>\<^sub>r ivl' \<and> ivl' \<le> snd (split_float_interval ivl) \<and> cmp ivly \<and> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
856 |
approx_tse prec 0 t (mid ivl') 1 f [Some ivl'] = Some ivly" |
60680 | 857 |
by blast |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
858 |
then show ?thesis |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
859 |
using subintervals |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
860 |
by (auto) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
861 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
862 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
863 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
864 |
lemma approx_tse_form'_less: |
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
865 |
fixes x :: real |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
866 |
assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s ivl (\<lambda> ivl. 0 < lower ivl)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
867 |
and x: "x \<in>\<^sub>r ivl" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
868 |
shows "interpret_floatarith b [x] < interpret_floatarith a [x]" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
869 |
proof - |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
870 |
from approx_tse_form'[OF tse x] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
871 |
obtain ivl' ivly |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
872 |
where x': "x \<in>\<^sub>r ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
873 |
and "ivl' \<le> ivl" and "0 < lower ivly" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
874 |
and tse: "approx_tse prec 0 t (mid ivl') 1 (Add a (Minus b)) [Some ivl'] = Some ivly" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
875 |
by blast |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
876 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
877 |
hence "bounded_by [x] [Some ivl']" |
60680 | 878 |
by (auto simp add: bounded_by_def) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
879 |
from approx_tse[OF this _ _ _ _ tse, of ivl'] x' mid_le |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
880 |
have "lower ivly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
881 |
by (auto simp: set_of_eq) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
882 |
from order_less_le_trans[OF _ this, of 0] \<open>0 < lower ivly\<close> show ?thesis |
60680 | 883 |
by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
884 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
885 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
886 |
lemma approx_tse_form'_le: |
40881
e84f82418e09
Use coercions in Approximation (by Dmitriy Traytel).
hoelzl
parents:
39556
diff
changeset
|
887 |
fixes x :: real |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
888 |
assumes tse: "approx_tse_form' prec t (Add a (Minus b)) s ivl (\<lambda> ivl. 0 \<le> lower ivl)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
889 |
and x: "x \<in>\<^sub>r ivl" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
890 |
shows "interpret_floatarith b [x] \<le> interpret_floatarith a [x]" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
891 |
proof - |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
892 |
from approx_tse_form'[OF tse x] |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
893 |
obtain ivl' ivly |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
894 |
where x': "x \<in>\<^sub>r ivl'" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
895 |
and "ivl' \<le> ivl" and "0 \<le> lower ivly" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
896 |
and tse: "approx_tse prec 0 t (mid ivl') 1 (Add a (Minus b)) [Some (ivl')] = Some ivly" |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
897 |
by blast |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
898 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
899 |
hence "bounded_by [x] [Some ivl']" by (auto simp add: bounded_by_def) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
900 |
from approx_tse[OF this _ _ _ _ tse, of ivl'] x' mid_le |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
901 |
have "lower ivly \<le> interpret_floatarith a [x] - interpret_floatarith b [x]" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
902 |
by (auto simp: set_of_eq) |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
903 |
from order_trans[OF _ this, of 0] \<open>0 \<le> lower ivly\<close> show ?thesis |
60680 | 904 |
by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
905 |
qed |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
906 |
|
58986 | 907 |
fun approx_tse_concl where |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
908 |
"approx_tse_concl prec t (Less lf rt) s ivl ivl' \<longleftrightarrow> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
909 |
approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 < lower ivl)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
910 |
"approx_tse_concl prec t (LessEqual lf rt) s ivl ivl' \<longleftrightarrow> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
911 |
approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
912 |
"approx_tse_concl prec t (AtLeastAtMost x lf rt) s ivl ivl' \<longleftrightarrow> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
913 |
(if approx_tse_form' prec t (Add x (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl) then |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
914 |
approx_tse_form' prec t (Add rt (Minus x)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl) else False)" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
915 |
"approx_tse_concl prec t (Conj f g) s ivl ivl' \<longleftrightarrow> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
916 |
approx_tse_concl prec t f s ivl ivl' \<and> approx_tse_concl prec t g s ivl ivl'" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
917 |
"approx_tse_concl prec t (Disj f g) s ivl ivl' \<longleftrightarrow> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
918 |
approx_tse_concl prec t f s ivl ivl' \<or> approx_tse_concl prec t g s ivl ivl'" | |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
919 |
"approx_tse_concl _ _ _ _ _ _ \<longleftrightarrow> False" |
58986 | 920 |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
921 |
definition |
60680 | 922 |
"approx_tse_form prec t s f = |
923 |
(case f of |
|
924 |
Bound x a b f \<Rightarrow> |
|
925 |
x = Var 0 \<and> |
|
926 |
(case (approx prec a [None], approx prec b [None]) of |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
927 |
(Some ivl, Some ivl') \<Rightarrow> approx_tse_concl prec t f s ivl ivl' |
60680 | 928 |
| _ \<Rightarrow> False) |
929 |
| _ \<Rightarrow> False)" |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
930 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
931 |
lemma approx_tse_form: |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
932 |
assumes "approx_tse_form prec t s f" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
933 |
shows "interpret_form f [x]" |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
934 |
proof (cases f) |
60680 | 935 |
case f_def: (Bound i a b f') |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
936 |
with assms obtain ivl ivl' |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
937 |
where a: "approx prec a [None] = Some ivl" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
938 |
and b: "approx prec b [None] = Some ivl'" |
55413
a8e96847523c
adapted theories to '{case,rec}_{list,option}' names
blanchet
parents:
54782
diff
changeset
|
939 |
unfolding approx_tse_form_def by (auto elim!: case_optionE) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
940 |
|
60680 | 941 |
from f_def assms have "i = Var 0" |
942 |
unfolding approx_tse_form_def by auto |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
943 |
hence i: "interpret_floatarith i [x] = x" by auto |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
944 |
|
60680 | 945 |
{ |
946 |
let ?f = "\<lambda>z. interpret_floatarith z [x]" |
|
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
947 |
assume "?f i \<in> { ?f a .. ?f b }" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
948 |
with approx[OF _ a, of "[x]"] approx[OF _ b, of "[x]"] |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
949 |
have bnd: "x \<in>\<^sub>r sup ivl ivl'" unfolding bounded_by_def i |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
950 |
by (auto simp: set_of_eq sup_float_def inf_float_def min_def max_def) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
951 |
|
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
952 |
have "interpret_form f' [x]" |
60680 | 953 |
using assms[unfolded f_def] |
58986 | 954 |
proof (induct f') |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
955 |
case (Less lf rt) |
58986 | 956 |
with a b |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
957 |
have "approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 < lower ivl)" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32920
diff
changeset
|
958 |
unfolding approx_tse_form_def by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
959 |
from approx_tse_form'_less[OF this bnd] |
58986 | 960 |
show ?case using Less by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
961 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
962 |
case (LessEqual lf rt) |
60680 | 963 |
with f_def a b assms |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
964 |
have "approx_tse_form' prec t (Add rt (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32920
diff
changeset
|
965 |
unfolding approx_tse_form_def by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
966 |
from approx_tse_form'_le[OF this bnd] |
58986 | 967 |
show ?case using LessEqual by auto |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
968 |
next |
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
969 |
case (AtLeastAtMost x lf rt) |
60680 | 970 |
with f_def a b assms |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
971 |
have "approx_tse_form' prec t (Add rt (Minus x)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
972 |
and "approx_tse_form' prec t (Add x (Minus lf)) s (sup ivl ivl') (\<lambda> ivl. 0 \<le> lower ivl)" |
62390 | 973 |
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
|
974 |
from approx_tse_form'_le[OF this(1) bnd] approx_tse_form'_le[OF this(2) bnd] |
58986 | 975 |
show ?case using AtLeastAtMost by auto |
976 |
qed (auto simp: f_def approx_tse_form_def elim!: case_optionE) |
|
60680 | 977 |
} |
978 |
thus ?thesis unfolding f_def by auto |
|
58986 | 979 |
qed (insert assms, auto simp add: approx_tse_form_def) |
31863
e391eee8bf14
Implemented taylor series expansion for approximation
hoelzl
parents:
31811
diff
changeset
|
980 |
|
69597 | 981 |
text \<open>\<^term>\<open>approx_form_eval\<close> 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
|
982 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
983 |
fun approx_form_eval :: "nat \<Rightarrow> form \<Rightarrow> (float interval) option list \<Rightarrow> (float interval) option list" where |
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
|
984 |
"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
|
985 |
(case (approx prec a bs, approx prec b bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
986 |
of (Some ivl1, Some ivl2) \<Rightarrow> approx_form_eval prec f (bs[n := Some (sup ivl1 ivl2)]) |
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
|
987 |
| _ \<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
|
988 |
"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
|
989 |
(case (approx prec a bs) |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
990 |
of (Some ivl) \<Rightarrow> approx_form_eval prec f (bs[n := Some ivl]) |
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
|
991 |
| _ \<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
|
992 |
"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
|
993 |
"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
|
994 |
"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
|
995 |
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
|
996 |
"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
|
997 |
|
60680 | 998 |
|
60533 | 999 |
subsection \<open>Implement proof method \texttt{approximation}\<close> |
29805 | 1000 |
|
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1001 |
ML \<open> |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1002 |
val _ = \<comment> \<open>Trusting the oracle \<close>@{oracle_name "holds_by_evaluation"} |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1003 |
signature APPROXIMATION_COMPUTATION = sig |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1004 |
val approx_bool: Proof.context -> term -> term |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1005 |
val approx_arith: Proof.context -> term -> term |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1006 |
val approx_form_eval: Proof.context -> term -> term |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1007 |
val approx_conv: Proof.context -> conv |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1008 |
end |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1009 |
|
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1010 |
structure Approximation_Computation : APPROXIMATION_COMPUTATION = struct |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1011 |
|
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1012 |
fun approx_conv ctxt ct = |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1013 |
@{computation_check |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1014 |
terms: Trueprop |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1015 |
"0 :: nat" "1 :: nat" "2 :: nat" "3 :: nat" Suc |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1016 |
"(+)::nat\<Rightarrow>nat\<Rightarrow>nat" "(-)::nat\<Rightarrow>nat\<Rightarrow>nat" "(*)::nat\<Rightarrow>nat\<Rightarrow>nat" |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1017 |
"0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1018 |
"(+)::int\<Rightarrow>int\<Rightarrow>int" "(-)::int\<Rightarrow>int\<Rightarrow>int" "(*)::int\<Rightarrow>int\<Rightarrow>int" "uminus::int\<Rightarrow>int" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1019 |
"replicate :: _ \<Rightarrow> (float interval) option \<Rightarrow> _" |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1020 |
"interval_of::float\<Rightarrow>float interval" |
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1021 |
approx' |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1022 |
approx_form |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1023 |
approx_tse_form |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1024 |
approx_form_eval |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1025 |
datatypes: bool |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1026 |
int |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1027 |
nat |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1028 |
integer |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1029 |
"nat list" |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1030 |
float |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1031 |
"(float interval) option list" |
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1032 |
floatarith |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1033 |
form |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1034 |
} ctxt ct |
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1035 |
|
74397 | 1036 |
fun term_of_bool true = \<^Const>\<open>True\<close> |
1037 |
| term_of_bool false = \<^Const>\<open>False\<close>; |
|
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1038 |
|
74397 | 1039 |
val mk_int = HOLogic.mk_number \<^Type>\<open>int\<close> o @{code integer_of_int}; |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
49962
diff
changeset
|
1040 |
|
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1041 |
fun term_of_float (@{code Float} (k, l)) = |
74397 | 1042 |
\<^Const>\<open>Float for \<open>mk_int k\<close> \<open>mk_int l\<close>\<close>; |
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1043 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1044 |
fun term_of_float_interval x = @{term "Interval::_\<Rightarrow>float interval"} $ |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1045 |
HOLogic.mk_prod |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1046 |
(apply2 term_of_float (@{code lowerF} x, @{code upperF} x)) |
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1047 |
|
74397 | 1048 |
fun term_of_float_interval_option NONE = \<^Const>\<open>None \<^typ>\<open>float interval option\<close>\<close> |
1049 |
| term_of_float_interval_option (SOME ff) = |
|
1050 |
\<^Const>\<open>Some \<^typ>\<open>float interval\<close> for \<open>term_of_float_interval ff\<close>\<close>; |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1051 |
|
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1052 |
val term_of_float_interval_option_list = |
74397 | 1053 |
HOLogic.mk_list \<^typ>\<open>float interval option\<close> o map term_of_float_interval_option; |
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1054 |
|
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1055 |
val approx_bool = @{computation bool} |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1056 |
(fn _ => fn x => case x of SOME b => term_of_bool b |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1057 |
| NONE => error "Computation approx_bool failed.") |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1058 |
val approx_arith = @{computation "float interval option"} |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1059 |
(fn _ => fn x => case x of SOME ffo => term_of_float_interval_option ffo |
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1060 |
| NONE => error "Computation approx_arith failed.") |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1061 |
val approx_form_eval = @{computation "float interval option list"} |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
1062 |
(fn _ => fn x => case x of SOME ffos => term_of_float_interval_option_list ffos |
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1063 |
| NONE => error "Computation approx_form_eval failed.") |
36985
41c5d4002f60
spelt out normalizer explicitly -- avoid dynamic reference to code generator configuration; avoid using old Codegen.eval_term
haftmann
parents:
36960
diff
changeset
|
1064 |
|
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70113
diff
changeset
|
1065 |
end |
60533 | 1066 |
\<close> |
31099
03314c427b34
optimized Approximation by precompiling approx_inequality
hoelzl
parents:
31098
diff
changeset
|
1067 |
|
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
1068 |
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
|
1069 |
by auto |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
1070 |
|
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
1071 |
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
|
1072 |
by auto |
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
1073 |
|
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1074 |
named_theorems approximation_preproc |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1075 |
|
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1076 |
lemma approximation_preproc_floatarith[approximation_preproc]: |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1077 |
"0 = real_of_float 0" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1078 |
"1 = real_of_float 1" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1079 |
"0 = Float 0 0" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1080 |
"1 = Float 1 0" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1081 |
"numeral a = Float (numeral a) 0" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1082 |
"numeral a = real_of_float (numeral a)" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1083 |
"x - y = x + - y" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1084 |
"x / y = x * inverse y" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1085 |
"ceiling x = - floor (- x)" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1086 |
"log x y = ln y * inverse (ln x)" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1087 |
"sin x = cos (pi / 2 - x)" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1088 |
"tan x = sin x / cos x" |
63931
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1089 |
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
|
1090 |
|
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1091 |
lemma approximation_preproc_int[approximation_preproc]: |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1092 |
"real_of_int 0 = real_of_float 0" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1093 |
"real_of_int 1 = real_of_float 1" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1094 |
"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
|
1095 |
"real_of_int (- i) = - real_of_int i" |
63931
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1096 |
"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
|
1097 |
"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
|
1098 |
"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
|
1099 |
"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
|
1100 |
"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
|
1101 |
"real_of_int (abs i) = abs (real_of_int i)" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1102 |
"real_of_int (i ^ n) = (real_of_int i) ^ n" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1103 |
"real_of_int (numeral a) = real_of_float (numeral a)" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1104 |
"i mod j = i - i div j * j" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1105 |
"i = j \<longleftrightarrow> real_of_int i = real_of_int j" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1106 |
"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
|
1107 |
"i < j \<longleftrightarrow> real_of_int i < real_of_int j" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1108 |
"i \<in> {j .. k} \<longleftrightarrow> real_of_int i \<in> {real_of_int j .. real_of_int k}" |
64246 | 1109 |
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
|
1110 |
|
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1111 |
lemma approximation_preproc_nat[approximation_preproc]: |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1112 |
"real 0 = real_of_float 0" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1113 |
"real 1 = real_of_float 1" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1114 |
"real (i + j) = real i + real j" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1115 |
"real (i - j) = max (real i - real j) 0" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1116 |
"real (i * j) = real i * real j" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1117 |
"real (i div j) = real_of_int (floor (real i / real j))" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1118 |
"real (min i j) = min (real i) (real j)" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1119 |
"real (max i j) = max (real i) (real j)" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1120 |
"real (i ^ n) = (real i) ^ n" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1121 |
"real (numeral a) = real_of_float (numeral a)" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1122 |
"i mod j = i - i div j * j" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1123 |
"n = m \<longleftrightarrow> real n = real m" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1124 |
"n \<le> m \<longleftrightarrow> real n \<le> real m" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1125 |
"n < m \<longleftrightarrow> real n < real m" |
f17a1c60ac39
approximation: preprocessing for nat/int expressions
immler
parents:
63929
diff
changeset
|
1126 |
"n \<in> {m .. l} \<longleftrightarrow> real n \<in> {real m .. real l}" |
64243 | 1127 |
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
|
1128 |
|
69605 | 1129 |
ML_file \<open>approximation.ML\<close> |
59850 | 1130 |
|
60533 | 1131 |
method_setup approximation = \<open> |
60680 | 1132 |
let |
1133 |
val free = |
|
1134 |
Args.context -- Args.term >> (fn (_, Free (n, _)) => n | (ctxt, t) => |
|
1135 |
error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); |
|
59850 | 1136 |
in |
60680 | 1137 |
Scan.lift Parse.nat -- |
59850 | 1138 |
Scan.optional (Scan.lift (Args.$$$ "splitting" |-- Args.colon) |
60680 | 1139 |
|-- Parse.and_list' (free --| Scan.lift (Args.$$$ "=") -- Scan.lift Parse.nat)) [] -- |
1140 |
Scan.option (Scan.lift (Args.$$$ "taylor" |-- Args.colon) |-- |
|
1141 |
(free |-- Scan.lift (Args.$$$ "=") |-- Scan.lift Parse.nat)) >> |
|
59850 | 1142 |
(fn ((prec, splitting), taylor) => fn ctxt => |
1143 |
SIMPLE_METHOD' (Approximation.approximation_tac prec splitting taylor ctxt)) |
|
1144 |
end |
|
60533 | 1145 |
\<close> "real number approximation" |
31811
64dea9a15031
Improved computation of bounds and implemented interval splitting for 'approximation'.
hoelzl
parents:
31810
diff
changeset
|
1146 |
|
58988 | 1147 |
|
1148 |
section "Quickcheck Generator" |
|
1149 |
||
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1150 |
lemma approximation_preproc_push_neg[approximation_preproc]: |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1151 |
fixes a b::real |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1152 |
shows |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1153 |
"\<not> (a < b) \<longleftrightarrow> b \<le> a" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1154 |
"\<not> (a \<le> b) \<longleftrightarrow> b < a" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1155 |
"\<not> (a = b) \<longleftrightarrow> b < a \<or> a < b" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1156 |
"\<not> (p \<and> q) \<longleftrightarrow> \<not> p \<or> \<not> q" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1157 |
"\<not> (p \<or> q) \<longleftrightarrow> \<not> p \<and> \<not> q" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1158 |
"\<not> \<not> q \<longleftrightarrow> q" |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1159 |
by auto |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
63918
diff
changeset
|
1160 |
|
69605 | 1161 |
ML_file \<open>approximation_generator.ML\<close> |
58988 | 1162 |
setup "Approximation_Generator.setup" |
1163 |
||
69830 | 1164 |
section "Avoid pollution of name space" |
1165 |
||
69835 | 1166 |
bundle floatarith_notation begin |
1167 |
||
1168 |
notation Add ("Add") |
|
1169 |
notation Minus ("Minus") |
|
1170 |
notation Mult ("Mult") |
|
1171 |
notation Inverse ("Inverse") |
|
1172 |
notation Cos ("Cos") |
|
1173 |
notation Arctan ("Arctan") |
|
1174 |
notation Abs ("Abs") |
|
1175 |
notation Max ("Max") |
|
1176 |
notation Min ("Min") |
|
1177 |
notation Pi ("Pi") |
|
1178 |
notation Sqrt ("Sqrt") |
|
1179 |
notation Exp ("Exp") |
|
1180 |
notation Powr ("Powr") |
|
1181 |
notation Ln ("Ln") |
|
1182 |
notation Power ("Power") |
|
1183 |
notation Floor ("Floor") |
|
1184 |
notation Var ("Var") |
|
1185 |
notation Num ("Num") |
|
1186 |
||
1187 |
end |
|
1188 |
||
1189 |
bundle no_floatarith_notation begin |
|
1190 |
||
1191 |
no_notation Add ("Add") |
|
1192 |
no_notation Minus ("Minus") |
|
1193 |
no_notation Mult ("Mult") |
|
1194 |
no_notation Inverse ("Inverse") |
|
1195 |
no_notation Cos ("Cos") |
|
1196 |
no_notation Arctan ("Arctan") |
|
1197 |
no_notation Abs ("Abs") |
|
1198 |
no_notation Max ("Max") |
|
1199 |
no_notation Min ("Min") |
|
1200 |
no_notation Pi ("Pi") |
|
1201 |
no_notation Sqrt ("Sqrt") |
|
1202 |
no_notation Exp ("Exp") |
|
1203 |
no_notation Powr ("Powr") |
|
1204 |
no_notation Ln ("Ln") |
|
1205 |
no_notation Power ("Power") |
|
1206 |
no_notation Floor ("Floor") |
|
1207 |
no_notation Var ("Var") |
|
1208 |
no_notation Num ("Num") |
|
1209 |
||
1210 |
end |
|
1211 |
||
69830 | 1212 |
hide_const (open) |
1213 |
Add |
|
1214 |
Minus |
|
1215 |
Mult |
|
1216 |
Inverse |
|
1217 |
Cos |
|
1218 |
Arctan |
|
1219 |
Abs |
|
1220 |
Max |
|
1221 |
Min |
|
1222 |
Pi |
|
1223 |
Sqrt |
|
1224 |
Exp |
|
1225 |
Powr |
|
1226 |
Ln |
|
1227 |
Power |
|
1228 |
Floor |
|
1229 |
Var |
|
1230 |
Num |
|
1231 |
||
29805 | 1232 |
end |