author | wenzelm |
Mon, 12 Oct 2015 19:47:23 +0200 | |
changeset 61410 | f569907de061 |
parent 60940 | 4c108cce6b35 |
child 62916 | 621afc4607ec |
permissions | -rw-r--r-- |
23178 | 1 |
(* Title: Pure/search.ML |
60940 | 2 |
Author: Lawrence C Paulson |
3 |
Author: Norbert Voelker, FernUniversitaet Hagen |
|
1588 | 4 |
|
16179 | 5 |
Search tacticals. |
1588 | 6 |
*) |
7 |
||
2672
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
8 |
infix 1 THEN_MAYBE THEN_MAYBE'; |
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
9 |
|
1588 | 10 |
signature SEARCH = |
32940 | 11 |
sig |
12 |
val trace_DEPTH_FIRST: bool Unsynchronized.ref |
|
13 |
val DEPTH_FIRST: (thm -> bool) -> tactic -> tactic |
|
14 |
val has_fewer_prems: int -> thm -> bool |
|
15 |
val IF_UNSOLVED: tactic -> tactic |
|
16 |
val SOLVE: tactic -> tactic |
|
17 |
val THEN_MAYBE: tactic * tactic -> tactic |
|
18 |
val THEN_MAYBE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic |
|
19 |
val DEPTH_SOLVE: tactic -> tactic |
|
20 |
val DEPTH_SOLVE_1: tactic -> tactic |
|
38802
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
21 |
val THEN_ITER_DEEPEN: int -> tactic -> (thm -> bool) -> (int -> tactic) -> tactic |
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
22 |
val ITER_DEEPEN: int -> (thm -> bool) -> (int -> tactic) -> tactic |
33335
1e189f96c393
DEEPEN: less aggressive tracing, subject to trace_DEEPEN;
wenzelm
parents:
32940
diff
changeset
|
23 |
val trace_DEEPEN: bool Unsynchronized.ref |
32940 | 24 |
val DEEPEN: int * int -> (int -> int -> tactic) -> int -> int -> tactic |
25 |
val trace_BEST_FIRST: bool Unsynchronized.ref |
|
26 |
val THEN_BEST_FIRST: tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic |
|
27 |
val BEST_FIRST: (thm -> bool) * (thm -> int) -> tactic -> tactic |
|
28 |
val BREADTH_FIRST: (thm -> bool) -> tactic -> tactic |
|
29 |
val QUIET_BREADTH_FIRST: (thm -> bool) -> tactic -> tactic |
|
30 |
val trace_ASTAR: bool Unsynchronized.ref |
|
31 |
val THEN_ASTAR: tactic -> (thm -> bool) * (int -> thm -> int) -> tactic -> tactic |
|
32 |
val ASTAR: (thm -> bool) * (int -> thm -> int) -> tactic -> tactic |
|
33 |
end; |
|
1588 | 34 |
|
32940 | 35 |
structure Search: SEARCH = |
1588 | 36 |
struct |
37 |
||
38 |
(**** Depth-first search ****) |
|
39 |
||
32738 | 40 |
val trace_DEPTH_FIRST = Unsynchronized.ref false; |
1588 | 41 |
|
42 |
(*Searches until "satp" reports proof tree as satisfied. |
|
43 |
Suppresses duplicate solutions to minimize search space.*) |
|
23178 | 44 |
fun DEPTH_FIRST satp tac = |
60940 | 45 |
let |
46 |
val tac = tracify trace_DEPTH_FIRST tac |
|
47 |
fun depth used [] = NONE |
|
48 |
| depth used (q :: qs) = |
|
49 |
(case Seq.pull q of |
|
50 |
NONE => depth used qs |
|
51 |
| SOME (st, stq) => |
|
52 |
if satp st andalso not (member Thm.eq_thm used st) then |
|
53 |
SOME (st, Seq.make (fn() => depth (st :: used) (stq :: qs))) |
|
54 |
else depth used (tac st :: stq :: qs)); |
|
55 |
in traced_tac (fn st => depth [] [Seq.single st]) end; |
|
1588 | 56 |
|
57 |
||
58 |
(*Predicate: Does the rule have fewer than n premises?*) |
|
59582 | 59 |
fun has_fewer_prems n rule = Thm.nprems_of rule < n; |
1588 | 60 |
|
61 |
(*Apply a tactic if subgoals remain, else do nothing.*) |
|
62 |
val IF_UNSOLVED = COND (has_fewer_prems 1) all_tac; |
|
63 |
||
5754 | 64 |
(*Force a tactic to solve its goal completely, otherwise fail *) |
65 |
fun SOLVE tac = tac THEN COND (has_fewer_prems 1) all_tac no_tac; |
|
66 |
||
2672
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
67 |
(*Execute tac1, but only execute tac2 if there are at least as many subgoals |
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
68 |
as before. This ensures that tac2 is only applied to an outcome of tac1.*) |
23178 | 69 |
fun (tac1 THEN_MAYBE tac2) st = |
59582 | 70 |
(tac1 THEN COND (has_fewer_prems (Thm.nprems_of st)) all_tac tac2) st; |
2672
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
71 |
|
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
72 |
fun (tac1 THEN_MAYBE' tac2) x = tac1 x THEN_MAYBE tac2 x; |
85d7e800d754
Replaced "flat" by the Basis Library function List.concat
paulson
parents:
2143
diff
changeset
|
73 |
|
1588 | 74 |
(*Tactical to reduce the number of premises by 1. |
75 |
If no subgoals then it must fail! *) |
|
3538 | 76 |
fun DEPTH_SOLVE_1 tac st = st |> |
59582 | 77 |
(case Thm.nprems_of st of |
78 |
0 => no_tac |
|
79 |
| n => DEPTH_FIRST (has_fewer_prems n) tac); |
|
1588 | 80 |
|
81 |
(*Uses depth-first search to solve ALL subgoals*) |
|
82 |
val DEPTH_SOLVE = DEPTH_FIRST (has_fewer_prems 1); |
|
83 |
||
84 |
||
85 |
||
2869
acee08856cc9
DEEPEN now takes an upper bound for terminating searches
paulson
parents:
2672
diff
changeset
|
86 |
(**** Iterative deepening with pruning ****) |
1588 | 87 |
|
88 |
fun has_vars (Var _) = true |
|
60940 | 89 |
| has_vars (Abs (_, _, t)) = has_vars t |
90 |
| has_vars (f $ t) = has_vars f orelse has_vars t |
|
1588 | 91 |
| has_vars _ = false; |
92 |
||
93 |
(*Counting of primitive inferences is APPROXIMATE, as the step tactic |
|
94 |
may perform >1 inference*) |
|
95 |
||
96 |
(*Pruning of rigid ancestor to prevent backtracking*) |
|
23178 | 97 |
fun prune (new as (k', np':int, rgd', stq), qs) = |
60940 | 98 |
let |
99 |
fun prune_aux (qs, []) = new :: qs |
|
100 |
| prune_aux (qs, (k, np, rgd, q) :: rqs) = |
|
101 |
if np' + 1 = np andalso rgd then |
|
102 |
(if !trace_DEPTH_FIRST then |
|
103 |
tracing ("Pruning " ^ |
|
104 |
string_of_int (1+length rqs) ^ " levels") |
|
105 |
else (); |
|
106 |
(*Use OLD k: zero-cost solution; see Stickel, p 365*) |
|
107 |
(k, np', rgd', stq) :: qs) |
|
108 |
else prune_aux ((k, np, rgd, q) :: qs, rqs) |
|
109 |
fun take ([], rqs) = ([], rqs) |
|
110 |
| take (arg as ((k, np, rgd, stq) :: qs, rqs)) = |
|
111 |
if np' < np then take (qs, (k, np, rgd, stq) :: rqs) else arg; |
|
112 |
in prune_aux (take (qs, [])) end; |
|
1588 | 113 |
|
114 |
||
115 |
(*Depth-first iterative deepening search for a state that satisfies satp |
|
116 |
tactic tac0 sets up the initial goal queue, while tac1 searches it. |
|
117 |
The solution sequence is redundant: the cutoff heuristic makes it impossible |
|
118 |
to suppress solutions arising from earlier searches, as the accumulated cost |
|
119 |
(k) can be wrong.*) |
|
38802
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
120 |
fun THEN_ITER_DEEPEN lim tac0 satp tac1 = traced_tac (fn st => |
60940 | 121 |
let |
122 |
val countr = Unsynchronized.ref 0 |
|
123 |
and tf = tracify trace_DEPTH_FIRST (tac1 1) |
|
124 |
and qs0 = tac0 st |
|
1588 | 125 |
(*bnd = depth bound; inc = estimate of increment required next*) |
60940 | 126 |
fun depth (bnd, inc) [] = |
38802
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
127 |
if bnd > lim then |
60940 | 128 |
(if !trace_DEPTH_FIRST then |
129 |
tracing (string_of_int (! countr) ^ |
|
130 |
" inferences so far. Giving up at " ^ string_of_int bnd) |
|
131 |
else (); |
|
132 |
NONE) |
|
22025
7c5896919eb8
Introduction of a bound on DEPTH_ITER, to prevent the meson method from running forever
paulson
parents:
20852
diff
changeset
|
133 |
else |
60940 | 134 |
(if !trace_DEPTH_FIRST then |
135 |
tracing (string_of_int (!countr) ^ |
|
136 |
" inferences so far. Searching to depth " ^ string_of_int bnd) |
|
137 |
else (); |
|
138 |
(*larger increments make it run slower for the hard problems*) |
|
139 |
depth (bnd + inc, 10)) [(0, 1, false, qs0)] |
|
140 |
| depth (bnd, inc) ((k, np, rgd, q) :: qs) = |
|
141 |
if k >= bnd then depth (bnd, inc) qs |
|
1588 | 142 |
else |
60940 | 143 |
(case |
144 |
(Unsynchronized.inc countr; |
|
145 |
if !trace_DEPTH_FIRST then |
|
146 |
tracing (string_of_int np ^ implode (map (fn _ => "*") qs)) |
|
147 |
else (); |
|
148 |
Seq.pull q) of |
|
149 |
NONE => depth (bnd, inc) qs |
|
150 |
| SOME (st, stq) => |
|
151 |
if satp st then (*solution!*) |
|
152 |
SOME(st, Seq.make (fn() => depth (bnd, inc) ((k, np, rgd, stq) :: qs))) |
|
153 |
else |
|
154 |
let |
|
155 |
val np' = Thm.nprems_of st; |
|
156 |
(*rgd' calculation assumes tactic operates on subgoal 1*) |
|
157 |
val rgd' = not (has_vars (hd (Thm.prems_of st))); |
|
158 |
val k' = k + np' - np + 1; (*difference in # of subgoals, +1*) |
|
159 |
in |
|
160 |
if k' + np' >= bnd then depth (bnd, Int.min (inc, k' + np' + 1 - bnd)) qs |
|
161 |
else if np' < np (*solved a subgoal; prune rigid ancestors*) |
|
162 |
then depth (bnd, inc) (prune ((k', np', rgd', tf st), (k, np, rgd, stq) :: qs)) |
|
163 |
else depth (bnd, inc) ((k', np', rgd', tf st) :: (k, np, rgd, stq) :: qs) |
|
164 |
end) |
|
165 |
in depth (0, 5) [] end); |
|
1588 | 166 |
|
38802
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
167 |
fun ITER_DEEPEN lim = THEN_ITER_DEEPEN lim all_tac; |
1588 | 168 |
|
169 |
||
2869
acee08856cc9
DEEPEN now takes an upper bound for terminating searches
paulson
parents:
2672
diff
changeset
|
170 |
(*Simple iterative deepening tactical. It merely "deepens" any search tactic |
acee08856cc9
DEEPEN now takes an upper bound for terminating searches
paulson
parents:
2672
diff
changeset
|
171 |
using increment "inc" up to limit "lim". *) |
33335
1e189f96c393
DEEPEN: less aggressive tracing, subject to trace_DEEPEN;
wenzelm
parents:
32940
diff
changeset
|
172 |
val trace_DEEPEN = Unsynchronized.ref false; |
1e189f96c393
DEEPEN: less aggressive tracing, subject to trace_DEEPEN;
wenzelm
parents:
32940
diff
changeset
|
173 |
|
38802
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
174 |
fun DEEPEN (inc, lim) tacf m i = |
33380
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
175 |
let |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
176 |
fun dpn m st = |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
177 |
st |> |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
178 |
(if has_fewer_prems i st then no_tac |
38802
a925c0ee42f7
clarified iter_deepen_limit vs meson (cf. 7c5896919eb8) -- eliminated global ref;
wenzelm
parents:
35408
diff
changeset
|
179 |
else if m > lim then |
33380
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
180 |
(if !trace_DEEPEN then tracing "Search depth limit exceeded: giving up" else (); |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
181 |
no_tac) |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
182 |
else |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
183 |
(if !trace_DEEPEN then tracing ("Search depth = " ^ string_of_int m) else (); |
cd6023a9a922
DEEPEN: all tracing is subject to trace_DEEPEN (NB: Proof General tends to "popup" tracing output);
wenzelm
parents:
33335
diff
changeset
|
184 |
tacf m i ORELSE dpn (m+inc))) |
2869
acee08856cc9
DEEPEN now takes an upper bound for terminating searches
paulson
parents:
2672
diff
changeset
|
185 |
in dpn m end; |
23178 | 186 |
|
32940 | 187 |
|
1588 | 188 |
(*** Best-first search ***) |
189 |
||
32940 | 190 |
(*total ordering on theorems, allowing duplicates to be found*) |
191 |
structure Thm_Heap = Heap |
|
192 |
( |
|
193 |
type elem = int * thm; |
|
59058
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents:
55171
diff
changeset
|
194 |
val ord = prod_ord int_ord (Term_Ord.term_ord o apply2 Thm.prop_of); |
32940 | 195 |
); |
196 |
||
32738 | 197 |
val trace_BEST_FIRST = Unsynchronized.ref false; |
1588 | 198 |
|
199 |
(*For creating output sequence*) |
|
60940 | 200 |
fun some_of_list [] = NONE |
201 |
| some_of_list (x :: l) = SOME (x, Seq.make (fn () => some_of_list l)); |
|
1588 | 202 |
|
9094 | 203 |
(*Check for and delete duplicate proof states*) |
32940 | 204 |
fun delete_all_min prf heap = |
205 |
if Thm_Heap.is_empty heap then heap |
|
206 |
else if Thm.eq_thm (prf, #2 (Thm_Heap.min heap)) |
|
207 |
then delete_all_min prf (Thm_Heap.delete_min heap) |
|
208 |
else heap; |
|
1588 | 209 |
|
210 |
(*Best-first search for a state that satisfies satp (incl initial state) |
|
211 |
Function sizef estimates size of problem remaining (smaller means better). |
|
212 |
tactic tac0 sets up the initial priority queue, while tac1 searches it. *) |
|
23178 | 213 |
fun THEN_BEST_FIRST tac0 (satp, sizef) tac1 = |
60940 | 214 |
let |
215 |
val tac = tracify trace_BEST_FIRST tac1; |
|
216 |
fun pairsize th = (sizef th, th); |
|
217 |
fun bfs (news, nprf_heap) = |
|
218 |
(case List.partition satp news of |
|
219 |
([], nonsats) => next (fold_rev Thm_Heap.insert (map pairsize nonsats) nprf_heap) |
|
220 |
| (sats, _) => some_of_list sats) |
|
221 |
and next nprf_heap = |
|
222 |
if Thm_Heap.is_empty nprf_heap then NONE |
|
223 |
else |
|
224 |
let |
|
225 |
val (n, prf) = Thm_Heap.min nprf_heap; |
|
226 |
val _ = |
|
227 |
if !trace_BEST_FIRST |
|
228 |
then tracing("state size = " ^ string_of_int n) |
|
229 |
else (); |
|
230 |
in |
|
231 |
bfs (Seq.list_of (tac prf), delete_all_min prf (Thm_Heap.delete_min nprf_heap)) |
|
232 |
end; |
|
233 |
fun btac st = bfs (Seq.list_of (tac0 st), Thm_Heap.empty) |
|
1588 | 234 |
in traced_tac btac end; |
235 |
||
236 |
(*Ordinary best-first search, with no initial tactic*) |
|
237 |
val BEST_FIRST = THEN_BEST_FIRST all_tac; |
|
238 |
||
23178 | 239 |
(*Breadth-first search to satisfy satpred (including initial state) |
1588 | 240 |
SLOW -- SHOULD NOT USE APPEND!*) |
60940 | 241 |
fun gen_BREADTH_FIRST message satpred (tac: tactic) = |
242 |
let |
|
243 |
val tacf = Seq.list_of o tac; |
|
244 |
fun bfs prfs = |
|
245 |
(case List.partition satpred prfs of |
|
246 |
([], []) => [] |
|
247 |
| ([], nonsats) => |
|
248 |
(message ("breadth=" ^ string_of_int (length nonsats)); |
|
249 |
bfs (maps tacf nonsats)) |
|
250 |
| (sats, _) => sats); |
|
251 |
in fn st => Seq.of_list (bfs [st]) end; |
|
1588 | 252 |
|
12262 | 253 |
val BREADTH_FIRST = gen_BREADTH_FIRST tracing; |
5693 | 254 |
val QUIET_BREADTH_FIRST = gen_BREADTH_FIRST (K ()); |
255 |
||
1588 | 256 |
|
60940 | 257 |
(* |
258 |
Implementation of A*-like proof procedure by modification |
|
259 |
of the existing code for BEST_FIRST and best_tac so that the |
|
260 |
current level of search is taken into account. |
|
23178 | 261 |
*) |
1588 | 262 |
|
60940 | 263 |
(*Insertion into priority queue of states, marked with level*) |
264 |
fun insert_with_level (lnth: int * int * thm) [] = [lnth] |
|
265 |
| insert_with_level (l, m, th) ((l', n, th') :: nths) = |
|
266 |
if n < m then (l', n, th') :: insert_with_level (l, m, th) nths |
|
267 |
else if n = m andalso Thm.eq_thm (th, th') then (l', n, th') :: nths |
|
268 |
else (l, m, th) :: (l', n, th') :: nths; |
|
1588 | 269 |
|
270 |
(*For creating output sequence*) |
|
60940 | 271 |
fun some_of_list [] = NONE |
272 |
| some_of_list (x :: xs) = SOME (x, Seq.make (fn () => some_of_list xs)); |
|
1588 | 273 |
|
32738 | 274 |
val trace_ASTAR = Unsynchronized.ref false; |
1588 | 275 |
|
23178 | 276 |
fun THEN_ASTAR tac0 (satp, costf) tac1 = |
60940 | 277 |
let |
278 |
val tf = tracify trace_ASTAR tac1; |
|
279 |
fun bfs (news, nprfs, level) = |
|
280 |
let fun cost thm = (level, costf level thm, thm) in |
|
281 |
(case List.partition satp news of |
|
282 |
([], nonsats) => next (fold_rev (insert_with_level o cost) nonsats nprfs) |
|
283 |
| (sats, _) => some_of_list sats) |
|
284 |
end |
|
285 |
and next [] = NONE |
|
286 |
| next ((level, n, prf) :: nprfs) = |
|
287 |
(if !trace_ASTAR then |
|
288 |
tracing ("level = " ^ string_of_int level ^ |
|
289 |
" cost = " ^ string_of_int n ^ |
|
290 |
" queue length =" ^ string_of_int (length nprfs)) |
|
291 |
else (); |
|
292 |
bfs (Seq.list_of (tf prf), nprfs, level + 1)) |
|
293 |
fun tf st = bfs (Seq.list_of (tac0 st), [], 0); |
|
1588 | 294 |
in traced_tac tf end; |
295 |
||
296 |
(*Ordinary ASTAR, with no initial tactic*) |
|
297 |
val ASTAR = THEN_ASTAR all_tac; |
|
298 |
||
299 |
end; |
|
300 |
||
301 |
open Search; |