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