author | wenzelm |
Mon, 12 Apr 2021 14:14:47 +0200 | |
changeset 73563 | 55b66a45bc94 |
parent 73019 | 05e2cab9af8b |
child 74282 | c2ee8d993d6a |
permissions | -rw-r--r-- |
56813 | 1 |
(* Title: HOL/Decision_Procs/approximation.ML |
2 |
Author: Johannes Hoelzl, TU Muenchen |
|
3 |
*) |
|
4 |
||
5 |
signature APPROXIMATION = |
|
6 |
sig |
|
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
7 |
val reify_form: Proof.context -> term -> term |
56813 | 8 |
val approx: int -> Proof.context -> term -> term |
59850 | 9 |
val approximate: Proof.context -> term -> term |
10 |
val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic |
|
56813 | 11 |
end |
12 |
||
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
13 |
structure Approximation = |
56813 | 14 |
struct |
15 |
||
60754 | 16 |
fun reorder_bounds_tac ctxt prems i = |
59850 | 17 |
let |
69597 | 18 |
fun variable_of_bound (Const (\<^const_name>\<open>Trueprop\<close>, _) $ |
19 |
(Const (\<^const_name>\<open>Set.member\<close>, _) $ |
|
59850 | 20 |
Free (name, _) $ _)) = name |
69597 | 21 |
| variable_of_bound (Const (\<^const_name>\<open>Trueprop\<close>, _) $ |
22 |
(Const (\<^const_name>\<open>HOL.eq\<close>, _) $ |
|
59850 | 23 |
Free (name, _) $ _)) = name |
24 |
| variable_of_bound t = raise TERM ("variable_of_bound", [t]) |
|
25 |
||
26 |
val variable_bounds |
|
27 |
= map (`(variable_of_bound o Thm.prop_of)) prems |
|
28 |
||
29 |
fun add_deps (name, bnds) |
|
30 |
= Graph.add_deps_acyclic (name, |
|
31 |
remove (op =) name (Term.add_free_names (Thm.prop_of bnds) [])) |
|
32 |
||
33 |
val order = Graph.empty |
|
34 |
|> fold Graph.new_node variable_bounds |
|
35 |
|> fold add_deps variable_bounds |
|
36 |
|> Graph.strong_conn |> map the_single |> rev |
|
37 |
|> map_filter (AList.lookup (op =) variable_bounds) |
|
38 |
||
60754 | 39 |
fun prepend_prem th tac = |
40 |
tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i |
|
59850 | 41 |
in |
42 |
fold prepend_prem order all_tac |
|
43 |
end |
|
44 |
||
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
45 |
fun approximate ctxt t = case fastype_of t |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
46 |
of \<^typ>\<open>bool\<close> => |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
47 |
Approximation_Computation.approx_bool ctxt t |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
48 |
| \<^typ>\<open>(float interval) option\<close> => |
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
49 |
Approximation_Computation.approx_arith ctxt t |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
50 |
| \<^typ>\<open>(float interval) option list\<close> => |
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
51 |
Approximation_Computation.approx_form_eval ctxt t |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
52 |
| _ => error ("Bad term: " ^ Syntax.string_of_term ctxt t); |
59850 | 53 |
|
54 |
fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let |
|
63930 | 55 |
fun lookup_splitting (Free (name, _)) = |
56 |
(case AList.lookup (op =) splitting name |
|
69597 | 57 |
of SOME s => HOLogic.mk_number \<^typ>\<open>nat\<close> s |
58 |
| NONE => \<^term>\<open>0 :: nat\<close>) |
|
63930 | 59 |
| lookup_splitting t = raise TERM ("lookup_splitting", [t]) |
59850 | 60 |
val vs = nth (Thm.prems_of st) (i - 1) |
61 |
|> Logic.strip_imp_concl |
|
62 |
|> HOLogic.dest_Trueprop |
|
63 |
|> Term.strip_comb |> snd |> List.last |
|
64 |
|> HOLogic.dest_list |
|
65 |
val p = prec |
|
69597 | 66 |
|> HOLogic.mk_number \<^typ>\<open>nat\<close> |
59850 | 67 |
|> Thm.cterm_of ctxt |
68 |
in case taylor |
|
69 |
of NONE => let |
|
70 |
val n = vs |> length |
|
69597 | 71 |
|> HOLogic.mk_number \<^typ>\<open>nat\<close> |
59850 | 72 |
|> Thm.cterm_of ctxt |
73 |
val s = vs |
|
74 |
|> map lookup_splitting |
|
69597 | 75 |
|> HOLogic.mk_list \<^typ>\<open>nat\<close> |
59850 | 76 |
|> Thm.cterm_of ctxt |
77 |
in |
|
69597 | 78 |
(resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), \<^typ>\<open>nat\<close>), n), |
79 |
((("prec", 0), \<^typ>\<open>nat\<close>), p), |
|
80 |
((("ss", 0), \<^typ>\<open>nat list\<close>), s)]) |
|
60754 | 81 |
@{thm approx_form}] i |
69597 | 82 |
THEN simp_tac (put_simpset (simpset_of \<^context>) ctxt) i) st |
59850 | 83 |
end |
84 |
||
85 |
| SOME t => |
|
86 |
if length vs <> 1 |
|
87 |
then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st])) |
|
88 |
else let |
|
89 |
val t = t |
|
69597 | 90 |
|> HOLogic.mk_number \<^typ>\<open>nat\<close> |
59850 | 91 |
|> Thm.cterm_of ctxt |
92 |
val s = vs |> map lookup_splitting |> hd |
|
93 |
|> Thm.cterm_of ctxt |
|
94 |
in |
|
69597 | 95 |
resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), \<^typ>\<open>nat\<close>), s), |
96 |
((("t", 0), \<^typ>\<open>nat\<close>), t), |
|
97 |
((("prec", 0), \<^typ>\<open>nat\<close>), p)]) |
|
60754 | 98 |
@{thm approx_tse_form}] i st |
59850 | 99 |
end |
100 |
end |
|
101 |
||
69597 | 102 |
fun calculated_subterms (\<^const>\<open>Trueprop\<close> $ t) = calculated_subterms t |
103 |
| calculated_subterms (\<^const>\<open>HOL.implies\<close> $ _ $ t) = calculated_subterms t |
|
104 |
| calculated_subterms (\<^term>\<open>(\<le>) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ t1 $ t2) = [t1, t2] |
|
105 |
| calculated_subterms (\<^term>\<open>(<) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ t1 $ t2) = [t1, t2] |
|
106 |
| calculated_subterms (\<^term>\<open>(\<in>) :: real \<Rightarrow> real set \<Rightarrow> bool\<close> $ t1 $ |
|
107 |
(\<^term>\<open>atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set\<close> $ t2 $ t3)) = [t1, t2, t3] |
|
56813 | 108 |
| calculated_subterms t = raise TERM ("calculated_subterms", [t]) |
109 |
||
69597 | 110 |
fun dest_interpret_form (\<^const>\<open>interpret_form\<close> $ b $ xs) = (b, xs) |
56813 | 111 |
| dest_interpret_form t = raise TERM ("dest_interpret_form", [t]) |
112 |
||
69597 | 113 |
fun dest_interpret (\<^const>\<open>interpret_floatarith\<close> $ b $ xs) = (b, xs) |
56813 | 114 |
| dest_interpret t = raise TERM ("dest_interpret", [t]) |
115 |
||
69597 | 116 |
fun dest_interpret_env (\<^const>\<open>interpret_form\<close> $ _ $ xs) = xs |
117 |
| dest_interpret_env (\<^const>\<open>interpret_floatarith\<close> $ _ $ xs) = xs |
|
63930 | 118 |
| dest_interpret_env t = raise TERM ("dest_interpret_env", [t]) |
56813 | 119 |
|
69597 | 120 |
fun dest_float (\<^const>\<open>Float\<close> $ m $ e) = (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e)) |
63930 | 121 |
| dest_float t = raise TERM ("dest_float", [t]) |
122 |
||
56813 | 123 |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
124 |
fun dest_ivl |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
125 |
(Const (\<^const_name>\<open>Some\<close>, _) $ |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
126 |
(Const (\<^const_name>\<open>Interval\<close>, _) $ ((Const (\<^const_name>\<open>Pair\<close>, _) $ u $ l)))) = |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
127 |
SOME (dest_float u, dest_float l) |
69597 | 128 |
| dest_ivl (Const (\<^const_name>\<open>None\<close>, _)) = NONE |
56813 | 129 |
| dest_ivl t = raise TERM ("dest_result", [t]) |
130 |
||
69597 | 131 |
fun mk_approx' prec t = (\<^const>\<open>approx'\<close> |
132 |
$ HOLogic.mk_number \<^typ>\<open>nat\<close> prec |
|
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
133 |
$ t $ \<^term>\<open>[] :: (float interval) option list\<close>) |
56813 | 134 |
|
69597 | 135 |
fun mk_approx_form_eval prec t xs = (\<^const>\<open>approx_form_eval\<close> |
136 |
$ HOLogic.mk_number \<^typ>\<open>nat\<close> prec |
|
56813 | 137 |
$ t $ xs) |
138 |
||
139 |
fun float2_float10 prec round_down (m, e) = ( |
|
140 |
let |
|
141 |
val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0)) |
|
142 |
||
63930 | 143 |
fun frac _ _ 0 digits cnt = (digits, cnt, 0) |
144 |
| frac _ 0 r digits cnt = (digits, cnt, r) |
|
56813 | 145 |
| frac c p r digits cnt = (let |
146 |
val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2) |
|
147 |
in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r |
|
148 |
(digits * 10 + d) (cnt + 1) |
|
149 |
end) |
|
150 |
||
151 |
val sgn = Int.sign m |
|
152 |
val m = abs m |
|
153 |
||
154 |
val round_down = (sgn = 1 andalso round_down) orelse |
|
155 |
(sgn = ~1 andalso not round_down) |
|
156 |
||
157 |
val (x, r) = Integer.div_mod m (Integer.pow (~e) 2) |
|
158 |
||
73019
05e2cab9af8b
tuned signature -- prefer Isabelle/ML structure Integer;
wenzelm
parents:
71037
diff
changeset
|
159 |
val p = ((if x = 0 then prec else prec - (Integer.log2 x + 1)) * 3) div 10 + 1 |
56813 | 160 |
|
161 |
val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0) |
|
162 |
||
163 |
val digits = if round_down orelse r = 0 then digits else digits + 1 |
|
164 |
||
165 |
in (sgn * (digits + x * (Integer.pow e10 10)), ~e10) |
|
166 |
end) |
|
167 |
||
168 |
fun mk_result prec (SOME (l, u)) = |
|
169 |
(let |
|
170 |
fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x |
|
69597 | 171 |
in if e = 0 then HOLogic.mk_number \<^typ>\<open>real\<close> m |
172 |
else if e = 1 then \<^term>\<open>divide :: real \<Rightarrow> real \<Rightarrow> real\<close> $ |
|
173 |
HOLogic.mk_number \<^typ>\<open>real\<close> m $ |
|
174 |
\<^term>\<open>10\<close> |
|
175 |
else \<^term>\<open>divide :: real \<Rightarrow> real \<Rightarrow> real\<close> $ |
|
176 |
HOLogic.mk_number \<^typ>\<open>real\<close> m $ |
|
177 |
(\<^term>\<open>power 10 :: nat \<Rightarrow> real\<close> $ |
|
178 |
HOLogic.mk_number \<^typ>\<open>nat\<close> (~e)) end) |
|
179 |
in \<^term>\<open>atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set\<close> $ mk_float10 true l $ mk_float10 false u end) |
|
180 |
| mk_result _ NONE = \<^term>\<open>UNIV :: real set\<close> |
|
56813 | 181 |
|
182 |
fun realify t = |
|
183 |
let |
|
184 |
val t = Logic.varify_global t |
|
69597 | 185 |
val m = map (fn (name, _) => (name, \<^typ>\<open>real\<close>)) (Term.add_tvars t []) |
56813 | 186 |
val t = Term.subst_TVars m t |
187 |
in t end |
|
188 |
||
189 |
fun apply_tactic ctxt term tactic = |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
190 |
Thm.cterm_of ctxt term |
56813 | 191 |
|> Goal.init |
192 |
|> SINGLE tactic |
|
59582 | 193 |
|> the |> Thm.prems_of |> hd |
56813 | 194 |
|
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
195 |
fun preproc_form_conv ctxt = |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
196 |
Simplifier.rewrite |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
197 |
(put_simpset HOL_basic_ss ctxt addsimps |
69597 | 198 |
(Named_Theorems.get ctxt \<^named_theorems>\<open>approximation_preproc\<close>)) |
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
199 |
|
63930 | 200 |
fun reify_form_conv ctxt ct = |
201 |
let |
|
202 |
val thm = |
|
203 |
Reification.conv ctxt @{thms interpret_form.simps interpret_floatarith.simps} ct |
|
204 |
handle ERROR msg => |
|
205 |
cat_error ("Reification failed: " ^ msg) |
|
206 |
("Approximation does not support " ^ |
|
207 |
quote (Syntax.string_of_term ctxt (Thm.term_of ct))) |
|
208 |
fun check_env (Free _) = () |
|
209 |
| check_env (Var _) = () |
|
210 |
| check_env t = |
|
211 |
cat_error "Term not supported by approximation:" (Syntax.string_of_term ctxt t) |
|
212 |
val _ = Thm.rhs_of thm |> Thm.term_of |> dest_interpret_env |> HOLogic.dest_list |> map check_env |
|
213 |
in thm end |
|
214 |
||
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
215 |
|
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
216 |
fun reify_form_tac ctxt i = CONVERSION (Conv.arg_conv (reify_form_conv ctxt)) i |
56813 | 217 |
|
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
218 |
fun prepare_form_tac ctxt i = |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
219 |
REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE}, |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
220 |
eresolve_tac ctxt @{thms meta_eqE}, |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
221 |
resolve_tac ctxt @{thms impI}] i) |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
222 |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt i |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
223 |
THEN DETERM (TRY (filter_prems_tac ctxt (K false) i)) |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
224 |
THEN CONVERSION (Conv.arg_conv (preproc_form_conv ctxt)) i |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
225 |
|
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
226 |
fun prepare_form ctxt term = apply_tactic ctxt term (prepare_form_tac ctxt 1) |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
227 |
|
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
228 |
fun apply_reify_form ctxt t = apply_tactic ctxt t (reify_form_tac ctxt 1) |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
229 |
|
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
230 |
fun reify_form ctxt t = HOLogic.mk_Trueprop t |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
231 |
|> prepare_form ctxt |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
232 |
|> apply_reify_form ctxt |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
233 |
|> HOLogic.dest_Trueprop |
56813 | 234 |
|
235 |
fun approx_form prec ctxt t = |
|
236 |
realify t |
|
237 |
|> prepare_form ctxt |
|
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
238 |
|> (fn arith_term => apply_reify_form ctxt arith_term |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
239 |
|> HOLogic.dest_Trueprop |
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
240 |
|> dest_interpret_form |
56813 | 241 |
|> (fn (data, xs) => |
71037
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
242 |
mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\<open>(float interval) option\<close> |
f630f2e707a6
refactor Approximation.thy to use more abstract type of intervals
immler
parents:
71034
diff
changeset
|
243 |
(map (fn _ => \<^term>\<open>None :: (float interval) option\<close>) (HOLogic.dest_list xs))) |
56813 | 244 |
|> approximate ctxt |
245 |
|> HOLogic.dest_list |
|
246 |
|> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term) |
|
69597 | 247 |
|> map (fn (elem, s) => \<^term>\<open>(\<in>) :: real \<Rightarrow> real set \<Rightarrow> bool\<close> $ elem $ mk_result prec (dest_ivl s)) |
56813 | 248 |
|> foldr1 HOLogic.mk_conj)) |
249 |
||
250 |
fun approx_arith prec ctxt t = realify t |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
251 |
|> Thm.cterm_of ctxt |
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
252 |
|> (preproc_form_conv ctxt then_conv reify_form_conv ctxt) |
59582 | 253 |
|> Thm.prop_of |
56813 | 254 |
|> Logic.dest_equals |> snd |
255 |
|> dest_interpret |> fst |
|
256 |
|> mk_approx' prec |
|
257 |
|> approximate ctxt |
|
258 |
|> dest_ivl |
|
259 |
|> mk_result prec |
|
260 |
||
261 |
fun approx prec ctxt t = |
|
69597 | 262 |
if type_of t = \<^typ>\<open>prop\<close> then approx_form prec ctxt t |
263 |
else if type_of t = \<^typ>\<open>bool\<close> then approx_form prec ctxt (\<^const>\<open>Trueprop\<close> $ t) |
|
56813 | 264 |
else approx_arith prec ctxt t |
265 |
||
56923 | 266 |
fun approximate_cmd modes raw_t state = |
267 |
let |
|
268 |
val ctxt = Toplevel.context_of state; |
|
269 |
val t = Syntax.read_term ctxt raw_t; |
|
270 |
val t' = approx 30 ctxt t; |
|
271 |
val ty' = Term.type_of t'; |
|
70308 | 272 |
val ctxt' = Proof_Context.augment t' ctxt; |
59174 | 273 |
in |
274 |
Print_Mode.with_modes modes (fn () => |
|
56923 | 275 |
Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, |
59174 | 276 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) () |
277 |
end |> Pretty.writeln; |
|
56923 | 278 |
|
279 |
val opt_modes = |
|
69597 | 280 |
Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) []; |
56923 | 281 |
|
282 |
val _ = |
|
69597 | 283 |
Outer_Syntax.command \<^command_keyword>\<open>approximate\<close> "print approximation of term" |
56923 | 284 |
(opt_modes -- Parse.term |
285 |
>> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t))); |
|
286 |
||
71034
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
287 |
fun approximation_tac prec splitting taylor ctxt = |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
288 |
prepare_form_tac ctxt |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
289 |
THEN' reify_form_tac ctxt |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
290 |
THEN' rewrite_interpret_form_tac ctxt prec splitting taylor |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
291 |
THEN' CONVERSION (Approximation_Computation.approx_conv ctxt) |
e0755162093f
replace approximation oracle by less ad-hoc @{computation}s
immler
parents:
70308
diff
changeset
|
292 |
THEN' resolve_tac ctxt [TrueI] |
63929
b673e7221b16
approximation: rewrite for reduction to base expressions
immler
parents:
62969
diff
changeset
|
293 |
|
62391 | 294 |
end; |