| author | haftmann | 
| Sat, 19 Oct 2019 09:15:41 +0000 | |
| changeset 70903 | c550368a4e29 | 
| parent 70308 | 7f568724d67e | 
| child 71034 | e0755162093f | 
| 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  | 
||
13  | 
structure Approximation: APPROXIMATION =  | 
|
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  | 
||
45  | 
fun approximation_conv ctxt ct =  | 
|
| 64547 | 46  | 
approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct);  | 
| 59850 | 47  | 
|
48  | 
fun approximate ctxt t =  | 
|
49  | 
approximation_oracle (Proof_Context.theory_of ctxt, t)  | 
|
50  | 
|> Thm.prop_of |> Logic.dest_equals |> snd;  | 
|
51  | 
||
52  | 
(* Should be in HOL.thy ? *)  | 
|
| 60754 | 53  | 
fun gen_eval_tac conv ctxt =  | 
54  | 
CONVERSION  | 
|
55  | 
(Object_Logic.judgment_conv ctxt (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt))  | 
|
56  | 
THEN' resolve_tac ctxt [TrueI]  | 
|
| 59850 | 57  | 
|
58  | 
fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let  | 
|
| 63930 | 59  | 
fun lookup_splitting (Free (name, _)) =  | 
60  | 
(case AList.lookup (op =) splitting name  | 
|
| 69597 | 61  | 
of SOME s => HOLogic.mk_number \<^typ>\<open>nat\<close> s  | 
62  | 
| NONE => \<^term>\<open>0 :: nat\<close>)  | 
|
| 63930 | 63  | 
      | lookup_splitting t = raise TERM ("lookup_splitting", [t])
 | 
| 59850 | 64  | 
val vs = nth (Thm.prems_of st) (i - 1)  | 
65  | 
|> Logic.strip_imp_concl  | 
|
66  | 
|> HOLogic.dest_Trueprop  | 
|
67  | 
|> Term.strip_comb |> snd |> List.last  | 
|
68  | 
|> HOLogic.dest_list  | 
|
69  | 
val p = prec  | 
|
| 69597 | 70  | 
|> HOLogic.mk_number \<^typ>\<open>nat\<close>  | 
| 59850 | 71  | 
|> Thm.cterm_of ctxt  | 
72  | 
in case taylor  | 
|
73  | 
of NONE => let  | 
|
74  | 
val n = vs |> length  | 
|
| 69597 | 75  | 
|> HOLogic.mk_number \<^typ>\<open>nat\<close>  | 
| 59850 | 76  | 
|> Thm.cterm_of ctxt  | 
77  | 
val s = vs  | 
|
78  | 
|> map lookup_splitting  | 
|
| 69597 | 79  | 
|> HOLogic.mk_list \<^typ>\<open>nat\<close>  | 
| 59850 | 80  | 
|> Thm.cterm_of ctxt  | 
81  | 
in  | 
|
| 69597 | 82  | 
       (resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), \<^typ>\<open>nat\<close>), n),
 | 
83  | 
                                   ((("prec", 0), \<^typ>\<open>nat\<close>), p),
 | 
|
84  | 
                                   ((("ss", 0), \<^typ>\<open>nat list\<close>), s)])
 | 
|
| 60754 | 85  | 
            @{thm approx_form}] i
 | 
| 69597 | 86  | 
THEN simp_tac (put_simpset (simpset_of \<^context>) ctxt) i) st  | 
| 59850 | 87  | 
end  | 
88  | 
||
89  | 
| SOME t =>  | 
|
90  | 
if length vs <> 1  | 
|
91  | 
     then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st]))
 | 
|
92  | 
else let  | 
|
93  | 
val t = t  | 
|
| 69597 | 94  | 
|> HOLogic.mk_number \<^typ>\<open>nat\<close>  | 
| 59850 | 95  | 
|> Thm.cterm_of ctxt  | 
96  | 
val s = vs |> map lookup_splitting |> hd  | 
|
97  | 
|> Thm.cterm_of ctxt  | 
|
98  | 
in  | 
|
| 69597 | 99  | 
       resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), \<^typ>\<open>nat\<close>), s),
 | 
100  | 
                                   ((("t", 0), \<^typ>\<open>nat\<close>), t),
 | 
|
101  | 
                                   ((("prec", 0), \<^typ>\<open>nat\<close>), p)])
 | 
|
| 60754 | 102  | 
            @{thm approx_tse_form}] i st
 | 
| 59850 | 103  | 
end  | 
104  | 
end  | 
|
105  | 
||
| 69597 | 106  | 
fun calculated_subterms (\<^const>\<open>Trueprop\<close> $ t) = calculated_subterms t  | 
107  | 
| calculated_subterms (\<^const>\<open>HOL.implies\<close> $ _ $ t) = calculated_subterms t  | 
|
108  | 
| calculated_subterms (\<^term>\<open>(\<le>) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ t1 $ t2) = [t1, t2]  | 
|
109  | 
| calculated_subterms (\<^term>\<open>(<) :: real \<Rightarrow> real \<Rightarrow> bool\<close> $ t1 $ t2) = [t1, t2]  | 
|
110  | 
| calculated_subterms (\<^term>\<open>(\<in>) :: real \<Rightarrow> real set \<Rightarrow> bool\<close> $ t1 $  | 
|
111  | 
(\<^term>\<open>atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set\<close> $ t2 $ t3)) = [t1, t2, t3]  | 
|
| 56813 | 112  | 
  | calculated_subterms t = raise TERM ("calculated_subterms", [t])
 | 
113  | 
||
| 69597 | 114  | 
fun dest_interpret_form (\<^const>\<open>interpret_form\<close> $ b $ xs) = (b, xs)  | 
| 56813 | 115  | 
  | dest_interpret_form t = raise TERM ("dest_interpret_form", [t])
 | 
116  | 
||
| 69597 | 117  | 
fun dest_interpret (\<^const>\<open>interpret_floatarith\<close> $ b $ xs) = (b, xs)  | 
| 56813 | 118  | 
  | dest_interpret t = raise TERM ("dest_interpret", [t])
 | 
119  | 
||
| 69597 | 120  | 
fun dest_interpret_env (\<^const>\<open>interpret_form\<close> $ _ $ xs) = xs  | 
121  | 
| dest_interpret_env (\<^const>\<open>interpret_floatarith\<close> $ _ $ xs) = xs  | 
|
| 63930 | 122  | 
  | dest_interpret_env t = raise TERM ("dest_interpret_env", [t])
 | 
| 56813 | 123  | 
|
| 69597 | 124  | 
fun dest_float (\<^const>\<open>Float\<close> $ m $ e) = (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e))  | 
| 63930 | 125  | 
  | dest_float t = raise TERM ("dest_float", [t])
 | 
126  | 
||
| 56813 | 127  | 
|
| 69597 | 128  | 
fun dest_ivl (Const (\<^const_name>\<open>Some\<close>, _) $  | 
129  | 
(Const (\<^const_name>\<open>Pair\<close>, _) $ u $ l)) = SOME (dest_float u, dest_float l)  | 
|
130  | 
| dest_ivl (Const (\<^const_name>\<open>None\<close>, _)) = NONE  | 
|
| 56813 | 131  | 
  | dest_ivl t = raise TERM ("dest_result", [t])
 | 
132  | 
||
| 69597 | 133  | 
fun mk_approx' prec t = (\<^const>\<open>approx'\<close>  | 
134  | 
$ HOLogic.mk_number \<^typ>\<open>nat\<close> prec  | 
|
135  | 
$ t $ \<^term>\<open>[] :: (float * float) option list\<close>)  | 
|
| 56813 | 136  | 
|
| 69597 | 137  | 
fun mk_approx_form_eval prec t xs = (\<^const>\<open>approx_form_eval\<close>  | 
138  | 
$ HOLogic.mk_number \<^typ>\<open>nat\<close> prec  | 
|
| 56813 | 139  | 
$ t $ xs)  | 
140  | 
||
141  | 
fun float2_float10 prec round_down (m, e) = (  | 
|
142  | 
let  | 
|
143  | 
val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0))  | 
|
144  | 
||
| 63930 | 145  | 
fun frac _ _ 0 digits cnt = (digits, cnt, 0)  | 
146  | 
| frac _ 0 r digits cnt = (digits, cnt, r)  | 
|
| 56813 | 147  | 
| frac c p r digits cnt = (let  | 
148  | 
val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2)  | 
|
149  | 
in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r  | 
|
150  | 
(digits * 10 + d) (cnt + 1)  | 
|
151  | 
end)  | 
|
152  | 
||
153  | 
val sgn = Int.sign m  | 
|
154  | 
val m = abs m  | 
|
155  | 
||
156  | 
val round_down = (sgn = 1 andalso round_down) orelse  | 
|
157  | 
(sgn = ~1 andalso not round_down)  | 
|
158  | 
||
159  | 
val (x, r) = Integer.div_mod m (Integer.pow (~e) 2)  | 
|
160  | 
||
161  | 
val p = ((if x = 0 then prec else prec - (IntInf.log2 x + 1)) * 3) div 10 + 1  | 
|
162  | 
||
163  | 
val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0)  | 
|
164  | 
||
165  | 
val digits = if round_down orelse r = 0 then digits else digits + 1  | 
|
166  | 
||
167  | 
in (sgn * (digits + x * (Integer.pow e10 10)), ~e10)  | 
|
168  | 
end)  | 
|
169  | 
||
170  | 
fun mk_result prec (SOME (l, u)) =  | 
|
171  | 
(let  | 
|
172  | 
fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x  | 
|
| 69597 | 173  | 
in if e = 0 then HOLogic.mk_number \<^typ>\<open>real\<close> m  | 
174  | 
else if e = 1 then \<^term>\<open>divide :: real \<Rightarrow> real \<Rightarrow> real\<close> $  | 
|
175  | 
HOLogic.mk_number \<^typ>\<open>real\<close> m $  | 
|
176  | 
\<^term>\<open>10\<close>  | 
|
177  | 
else \<^term>\<open>divide :: real \<Rightarrow> real \<Rightarrow> real\<close> $  | 
|
178  | 
HOLogic.mk_number \<^typ>\<open>real\<close> m $  | 
|
179  | 
(\<^term>\<open>power 10 :: nat \<Rightarrow> real\<close> $  | 
|
180  | 
HOLogic.mk_number \<^typ>\<open>nat\<close> (~e)) end)  | 
|
181  | 
in \<^term>\<open>atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set\<close> $ mk_float10 true l $ mk_float10 false u end)  | 
|
182  | 
| mk_result _ NONE = \<^term>\<open>UNIV :: real set\<close>  | 
|
| 56813 | 183  | 
|
184  | 
fun realify t =  | 
|
185  | 
let  | 
|
186  | 
val t = Logic.varify_global t  | 
|
| 69597 | 187  | 
val m = map (fn (name, _) => (name, \<^typ>\<open>real\<close>)) (Term.add_tvars t [])  | 
| 56813 | 188  | 
val t = Term.subst_TVars m t  | 
189  | 
in t end  | 
|
190  | 
||
191  | 
fun apply_tactic ctxt term tactic =  | 
|
| 
59621
 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 
wenzelm 
parents: 
59582 
diff
changeset
 | 
192  | 
Thm.cterm_of ctxt term  | 
| 56813 | 193  | 
|> Goal.init  | 
194  | 
|> SINGLE tactic  | 
|
| 59582 | 195  | 
|> the |> Thm.prems_of |> hd  | 
| 56813 | 196  | 
|
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
197  | 
fun preproc_form_conv ctxt =  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
198  | 
Simplifier.rewrite  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
199  | 
(put_simpset HOL_basic_ss ctxt addsimps  | 
| 69597 | 200  | 
(Named_Theorems.get ctxt \<^named_theorems>\<open>approximation_preproc\<close>))  | 
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
201  | 
|
| 63930 | 202  | 
fun reify_form_conv ctxt ct =  | 
203  | 
let  | 
|
204  | 
val thm =  | 
|
205  | 
       Reification.conv ctxt @{thms interpret_form.simps interpret_floatarith.simps} ct
 | 
|
206  | 
handle ERROR msg =>  | 
|
207  | 
        cat_error ("Reification failed: " ^ msg)
 | 
|
208  | 
          ("Approximation does not support " ^
 | 
|
209  | 
quote (Syntax.string_of_term ctxt (Thm.term_of ct)))  | 
|
210  | 
fun check_env (Free _) = ()  | 
|
211  | 
| check_env (Var _) = ()  | 
|
212  | 
| check_env t =  | 
|
213  | 
cat_error "Term not supported by approximation:" (Syntax.string_of_term ctxt t)  | 
|
214  | 
val _ = Thm.rhs_of thm |> Thm.term_of |> dest_interpret_env |> HOLogic.dest_list |> map check_env  | 
|
215  | 
in thm end  | 
|
216  | 
||
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
217  | 
|
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
218  | 
fun reify_form_tac ctxt i = CONVERSION (Conv.arg_conv (reify_form_conv ctxt)) i  | 
| 56813 | 219  | 
|
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
220  | 
fun prepare_form_tac ctxt i =  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
221  | 
  REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE},
 | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
222  | 
    eresolve_tac ctxt @{thms meta_eqE},
 | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
223  | 
    resolve_tac ctxt @{thms impI}] i)
 | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
224  | 
  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
 | 
225  | 
THEN DETERM (TRY (filter_prems_tac ctxt (K false) i))  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
226  | 
THEN CONVERSION (Conv.arg_conv (preproc_form_conv ctxt)) i  | 
| 
 
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 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
 | 
229  | 
|
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
230  | 
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
 | 
231  | 
|
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
232  | 
fun reify_form ctxt t = HOLogic.mk_Trueprop t  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
233  | 
|> prepare_form ctxt  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
234  | 
|> apply_reify_form ctxt  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
235  | 
|> HOLogic.dest_Trueprop  | 
| 56813 | 236  | 
|
237  | 
fun approx_form prec ctxt t =  | 
|
238  | 
realify t  | 
|
239  | 
|> prepare_form ctxt  | 
|
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
240  | 
|> (fn arith_term => apply_reify_form ctxt arith_term  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
241  | 
|> HOLogic.dest_Trueprop  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
242  | 
|> dest_interpret_form  | 
| 56813 | 243  | 
|> (fn (data, xs) =>  | 
| 69597 | 244  | 
mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\<open>(float * float) option\<close>  | 
245  | 
(map (fn _ => \<^term>\<open>None :: (float * float) option\<close>) (HOLogic.dest_list xs)))  | 
|
| 56813 | 246  | 
|> approximate ctxt  | 
247  | 
|> HOLogic.dest_list  | 
|
248  | 
|> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term)  | 
|
| 69597 | 249  | 
|> map (fn (elem, s) => \<^term>\<open>(\<in>) :: real \<Rightarrow> real set \<Rightarrow> bool\<close> $ elem $ mk_result prec (dest_ivl s))  | 
| 56813 | 250  | 
|> foldr1 HOLogic.mk_conj))  | 
251  | 
||
252  | 
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
 | 
253  | 
|> Thm.cterm_of ctxt  | 
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
254  | 
|> (preproc_form_conv ctxt then_conv reify_form_conv ctxt)  | 
| 59582 | 255  | 
|> Thm.prop_of  | 
| 56813 | 256  | 
|> Logic.dest_equals |> snd  | 
257  | 
|> dest_interpret |> fst  | 
|
258  | 
|> mk_approx' prec  | 
|
259  | 
|> approximate ctxt  | 
|
260  | 
|> dest_ivl  | 
|
261  | 
|> mk_result prec  | 
|
262  | 
||
263  | 
fun approx prec ctxt t =  | 
|
| 69597 | 264  | 
if type_of t = \<^typ>\<open>prop\<close> then approx_form prec ctxt t  | 
265  | 
else if type_of t = \<^typ>\<open>bool\<close> then approx_form prec ctxt (\<^const>\<open>Trueprop\<close> $ t)  | 
|
| 56813 | 266  | 
else approx_arith prec ctxt t  | 
267  | 
||
| 56923 | 268  | 
fun approximate_cmd modes raw_t state =  | 
269  | 
let  | 
|
270  | 
val ctxt = Toplevel.context_of state;  | 
|
271  | 
val t = Syntax.read_term ctxt raw_t;  | 
|
272  | 
val t' = approx 30 ctxt t;  | 
|
273  | 
val ty' = Term.type_of t';  | 
|
| 70308 | 274  | 
val ctxt' = Proof_Context.augment t' ctxt;  | 
| 59174 | 275  | 
in  | 
276  | 
Print_Mode.with_modes modes (fn () =>  | 
|
| 56923 | 277  | 
Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,  | 
| 59174 | 278  | 
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ()  | 
279  | 
end |> Pretty.writeln;  | 
|
| 56923 | 280  | 
|
281  | 
val opt_modes =  | 
|
| 69597 | 282  | 
Scan.optional (\<^keyword>\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\<open>)\<close>)) [];  | 
| 56923 | 283  | 
|
284  | 
val _ =  | 
|
| 69597 | 285  | 
Outer_Syntax.command \<^command_keyword>\<open>approximate\<close> "print approximation of term"  | 
| 56923 | 286  | 
(opt_modes -- Parse.term  | 
287  | 
>> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t)));  | 
|
288  | 
||
| 59850 | 289  | 
fun approximation_tac prec splitting taylor ctxt i =  | 
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
290  | 
prepare_form_tac ctxt i  | 
| 
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
291  | 
THEN reify_form_tac ctxt i  | 
| 59850 | 292  | 
THEN rewrite_interpret_form_tac ctxt prec splitting taylor i  | 
293  | 
THEN gen_eval_tac (approximation_conv ctxt) ctxt i  | 
|
| 
63929
 
b673e7221b16
approximation: rewrite for reduction to base expressions
 
immler 
parents: 
62969 
diff
changeset
 | 
294  | 
|
| 62391 | 295  | 
end;  |