author | blanchet |
Fri, 02 Oct 2015 21:06:32 +0200 | |
changeset 61310 | 9a50ea544fd3 |
parent 60754 | 02924903a6fd |
child 62391 | 1658fc9b2618 |
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 |
|
7 |
val approx: int -> Proof.context -> term -> term |
|
59850 | 8 |
val approximate: Proof.context -> term -> term |
9 |
val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic |
|
56813 | 10 |
end |
11 |
||
12 |
structure Approximation: APPROXIMATION = |
|
13 |
struct |
|
14 |
||
60754 | 15 |
fun reorder_bounds_tac ctxt prems i = |
59850 | 16 |
let |
17 |
fun variable_of_bound (Const (@{const_name Trueprop}, _) $ |
|
18 |
(Const (@{const_name Set.member}, _) $ |
|
19 |
Free (name, _) $ _)) = name |
|
20 |
| variable_of_bound (Const (@{const_name Trueprop}, _) $ |
|
21 |
(Const (@{const_name HOL.eq}, _) $ |
|
22 |
Free (name, _) $ _)) = name |
|
23 |
| variable_of_bound t = raise TERM ("variable_of_bound", [t]) |
|
24 |
||
25 |
val variable_bounds |
|
26 |
= map (`(variable_of_bound o Thm.prop_of)) prems |
|
27 |
||
28 |
fun add_deps (name, bnds) |
|
29 |
= Graph.add_deps_acyclic (name, |
|
30 |
remove (op =) name (Term.add_free_names (Thm.prop_of bnds) [])) |
|
31 |
||
32 |
val order = Graph.empty |
|
33 |
|> fold Graph.new_node variable_bounds |
|
34 |
|> fold add_deps variable_bounds |
|
35 |
|> Graph.strong_conn |> map the_single |> rev |
|
36 |
|> map_filter (AList.lookup (op =) variable_bounds) |
|
37 |
||
60754 | 38 |
fun prepend_prem th tac = |
39 |
tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i |
|
59850 | 40 |
in |
41 |
fold prepend_prem order all_tac |
|
42 |
end |
|
43 |
||
44 |
fun approximation_conv ctxt ct = |
|
45 |
approximation_oracle (Proof_Context.theory_of ctxt, Thm.term_of ct |> tap (tracing o Syntax.string_of_term ctxt)); |
|
46 |
||
47 |
fun approximate ctxt t = |
|
48 |
approximation_oracle (Proof_Context.theory_of ctxt, t) |
|
49 |
|> Thm.prop_of |> Logic.dest_equals |> snd; |
|
50 |
||
51 |
(* Should be in HOL.thy ? *) |
|
60754 | 52 |
fun gen_eval_tac conv ctxt = |
53 |
CONVERSION |
|
54 |
(Object_Logic.judgment_conv ctxt (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt)) |
|
55 |
THEN' resolve_tac ctxt [TrueI] |
|
59850 | 56 |
|
57 |
val form_equations = @{thms interpret_form_equations}; |
|
58 |
||
59 |
fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let |
|
60 |
fun lookup_splitting (Free (name, _)) |
|
61 |
= case AList.lookup (op =) splitting name |
|
62 |
of SOME s => HOLogic.mk_number @{typ nat} s |
|
63 |
| NONE => @{term "0 :: nat"} |
|
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 |
|
70 |
|> HOLogic.mk_number @{typ nat} |
|
71 |
|> Thm.cterm_of ctxt |
|
72 |
in case taylor |
|
73 |
of NONE => let |
|
74 |
val n = vs |> length |
|
75 |
|> HOLogic.mk_number @{typ nat} |
|
76 |
|> Thm.cterm_of ctxt |
|
77 |
val s = vs |
|
78 |
|> map lookup_splitting |
|
79 |
|> HOLogic.mk_list @{typ nat} |
|
80 |
|> Thm.cterm_of ctxt |
|
81 |
in |
|
60754 | 82 |
(resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), @{typ nat}), n), |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59970
diff
changeset
|
83 |
((("prec", 0), @{typ nat}), p), |
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59970
diff
changeset
|
84 |
((("ss", 0), @{typ "nat list"}), s)]) |
60754 | 85 |
@{thm approx_form}] i |
59850 | 86 |
THEN simp_tac (put_simpset (simpset_of @{context}) ctxt) i) st |
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 |
|
94 |
|> HOLogic.mk_number @{typ nat} |
|
95 |
|> Thm.cterm_of ctxt |
|
96 |
val s = vs |> map lookup_splitting |> hd |
|
97 |
|> Thm.cterm_of ctxt |
|
98 |
in |
|
60754 | 99 |
resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), @{typ nat}), s), |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59970
diff
changeset
|
100 |
((("t", 0), @{typ nat}), t), |
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59970
diff
changeset
|
101 |
((("prec", 0), @{typ nat}), p)]) |
60754 | 102 |
@{thm approx_tse_form}] i st |
59850 | 103 |
end |
104 |
end |
|
105 |
||
56813 | 106 |
fun calculated_subterms (@{const Trueprop} $ t) = calculated_subterms t |
107 |
| calculated_subterms (@{const HOL.implies} $ _ $ t) = calculated_subterms t |
|
108 |
| calculated_subterms (@{term "op <= :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2] |
|
109 |
| calculated_subterms (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = [t1, t2] |
|
110 |
| calculated_subterms (@{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ t1 $ |
|
111 |
(@{term "atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set"} $ t2 $ t3)) = [t1, t2, t3] |
|
112 |
| calculated_subterms t = raise TERM ("calculated_subterms", [t]) |
|
113 |
||
114 |
fun dest_interpret_form (@{const "interpret_form"} $ b $ xs) = (b, xs) |
|
115 |
| dest_interpret_form t = raise TERM ("dest_interpret_form", [t]) |
|
116 |
||
117 |
fun dest_interpret (@{const "interpret_floatarith"} $ b $ xs) = (b, xs) |
|
118 |
| dest_interpret t = raise TERM ("dest_interpret", [t]) |
|
119 |
||
120 |
||
121 |
fun dest_float (@{const "Float"} $ m $ e) = |
|
122 |
(snd (HOLogic.dest_number m), snd (HOLogic.dest_number e)) |
|
123 |
||
124 |
fun dest_ivl (Const (@{const_name "Some"}, _) $ |
|
125 |
(Const (@{const_name Pair}, _) $ u $ l)) = SOME (dest_float u, dest_float l) |
|
126 |
| dest_ivl (Const (@{const_name "None"}, _)) = NONE |
|
127 |
| dest_ivl t = raise TERM ("dest_result", [t]) |
|
128 |
||
129 |
fun mk_approx' prec t = (@{const "approx'"} |
|
130 |
$ HOLogic.mk_number @{typ nat} prec |
|
131 |
$ t $ @{term "[] :: (float * float) option list"}) |
|
132 |
||
133 |
fun mk_approx_form_eval prec t xs = (@{const "approx_form_eval"} |
|
134 |
$ HOLogic.mk_number @{typ nat} prec |
|
135 |
$ t $ xs) |
|
136 |
||
137 |
fun float2_float10 prec round_down (m, e) = ( |
|
138 |
let |
|
139 |
val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0)) |
|
140 |
||
141 |
fun frac c p 0 digits cnt = (digits, cnt, 0) |
|
142 |
| frac c 0 r digits cnt = (digits, cnt, r) |
|
143 |
| frac c p r digits cnt = (let |
|
144 |
val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2) |
|
145 |
in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r |
|
146 |
(digits * 10 + d) (cnt + 1) |
|
147 |
end) |
|
148 |
||
149 |
val sgn = Int.sign m |
|
150 |
val m = abs m |
|
151 |
||
152 |
val round_down = (sgn = 1 andalso round_down) orelse |
|
153 |
(sgn = ~1 andalso not round_down) |
|
154 |
||
155 |
val (x, r) = Integer.div_mod m (Integer.pow (~e) 2) |
|
156 |
||
157 |
val p = ((if x = 0 then prec else prec - (IntInf.log2 x + 1)) * 3) div 10 + 1 |
|
158 |
||
159 |
val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0) |
|
160 |
||
161 |
val digits = if round_down orelse r = 0 then digits else digits + 1 |
|
162 |
||
163 |
in (sgn * (digits + x * (Integer.pow e10 10)), ~e10) |
|
164 |
end) |
|
165 |
||
166 |
fun mk_result prec (SOME (l, u)) = |
|
167 |
(let |
|
168 |
fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x |
|
169 |
in if e = 0 then HOLogic.mk_number @{typ real} m |
|
170 |
else if e = 1 then @{term "divide :: real \<Rightarrow> real \<Rightarrow> real"} $ |
|
171 |
HOLogic.mk_number @{typ real} m $ |
|
172 |
@{term "10"} |
|
173 |
else @{term "divide :: real \<Rightarrow> real \<Rightarrow> real"} $ |
|
174 |
HOLogic.mk_number @{typ real} m $ |
|
175 |
(@{term "power 10 :: nat \<Rightarrow> real"} $ |
|
176 |
HOLogic.mk_number @{typ nat} (~e)) end) |
|
177 |
in @{term "atLeastAtMost :: real \<Rightarrow> real \<Rightarrow> real set"} $ mk_float10 true l $ mk_float10 false u end) |
|
178 |
| mk_result _ NONE = @{term "UNIV :: real set"} |
|
179 |
||
180 |
fun realify t = |
|
181 |
let |
|
182 |
val t = Logic.varify_global t |
|
183 |
val m = map (fn (name, _) => (name, @{typ real})) (Term.add_tvars t []) |
|
184 |
val t = Term.subst_TVars m t |
|
185 |
in t end |
|
186 |
||
187 |
fun apply_tactic ctxt term tactic = |
|
59621
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents:
59582
diff
changeset
|
188 |
Thm.cterm_of ctxt term |
56813 | 189 |
|> Goal.init |
190 |
|> SINGLE tactic |
|
59582 | 191 |
|> the |> Thm.prems_of |> hd |
56813 | 192 |
|
193 |
fun prepare_form ctxt term = apply_tactic ctxt term ( |
|
60754 | 194 |
REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE}, |
195 |
eresolve_tac ctxt @{thms meta_eqE}, |
|
196 |
resolve_tac ctxt @{thms impI}] 1) |
|
197 |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems 1) ctxt 1 |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59174
diff
changeset
|
198 |
THEN DETERM (TRY (filter_prems_tac ctxt (K false) 1))) |
56813 | 199 |
|
200 |
fun reify_form ctxt term = apply_tactic ctxt term |
|
201 |
(Reification.tac ctxt form_equations NONE 1) |
|
202 |
||
203 |
fun approx_form prec ctxt t = |
|
204 |
realify t |
|
205 |
|> prepare_form ctxt |
|
206 |
|> (fn arith_term => reify_form ctxt arith_term |
|
207 |
|> HOLogic.dest_Trueprop |> dest_interpret_form |
|
208 |
|> (fn (data, xs) => |
|
209 |
mk_approx_form_eval prec data (HOLogic.mk_list @{typ "(float * float) option"} |
|
210 |
(map (fn _ => @{term "None :: (float * float) option"}) (HOLogic.dest_list xs))) |
|
211 |
|> approximate ctxt |
|
212 |
|> HOLogic.dest_list |
|
213 |
|> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term) |
|
214 |
|> map (fn (elem, s) => @{term "op : :: real \<Rightarrow> real set \<Rightarrow> bool"} $ elem $ mk_result prec (dest_ivl s)) |
|
215 |
|> foldr1 HOLogic.mk_conj)) |
|
216 |
||
217 |
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
|
218 |
|> Thm.cterm_of ctxt |
56813 | 219 |
|> Reification.conv ctxt form_equations |
59582 | 220 |
|> Thm.prop_of |
56813 | 221 |
|> Logic.dest_equals |> snd |
222 |
|> dest_interpret |> fst |
|
223 |
|> mk_approx' prec |
|
224 |
|> approximate ctxt |
|
225 |
|> dest_ivl |
|
226 |
|> mk_result prec |
|
227 |
||
228 |
fun approx prec ctxt t = |
|
229 |
if type_of t = @{typ prop} then approx_form prec ctxt t |
|
230 |
else if type_of t = @{typ bool} then approx_form prec ctxt (@{const Trueprop} $ t) |
|
231 |
else approx_arith prec ctxt t |
|
232 |
||
56923 | 233 |
fun approximate_cmd modes raw_t state = |
234 |
let |
|
235 |
val ctxt = Toplevel.context_of state; |
|
236 |
val t = Syntax.read_term ctxt raw_t; |
|
237 |
val t' = approx 30 ctxt t; |
|
238 |
val ty' = Term.type_of t'; |
|
239 |
val ctxt' = Variable.auto_fixes t' ctxt; |
|
59174 | 240 |
in |
241 |
Print_Mode.with_modes modes (fn () => |
|
56923 | 242 |
Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, |
59174 | 243 |
Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) () |
244 |
end |> Pretty.writeln; |
|
56923 | 245 |
|
246 |
val opt_modes = |
|
247 |
Scan.optional (@{keyword "("} |-- Parse.!!! (Scan.repeat1 Parse.xname --| @{keyword ")"})) []; |
|
248 |
||
249 |
val _ = |
|
59936
b8ffc3dc9e24
@{command_spec} is superseded by @{command_keyword};
wenzelm
parents:
59850
diff
changeset
|
250 |
Outer_Syntax.command @{command_keyword approximate} "print approximation of term" |
56923 | 251 |
(opt_modes -- Parse.term |
252 |
>> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t))); |
|
253 |
||
59850 | 254 |
fun approximation_tac prec splitting taylor ctxt i = |
60754 | 255 |
REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE}, |
256 |
eresolve_tac ctxt @{thms meta_eqE}, |
|
257 |
resolve_tac ctxt @{thms impI}] i) |
|
258 |
THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt i |
|
59850 | 259 |
THEN DETERM (TRY (filter_prems_tac ctxt (K false) i)) |
260 |
THEN DETERM (Reification.tac ctxt form_equations NONE i) |
|
261 |
THEN rewrite_interpret_form_tac ctxt prec splitting taylor i |
|
262 |
THEN gen_eval_tac (approximation_conv ctxt) ctxt i |
|
263 |
||
264 |
end; |