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