author | wenzelm |
Mon, 14 Apr 1997 10:28:21 +0200 | |
changeset 2947 | abca00c27841 |
parent 2924 | af506c35b4ed |
child 2952 | ea834e8fac1b |
permissions | -rw-r--r-- |
2894 | 1 |
(* Title: Provers/blast |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
5 |
||
6 |
Generic tableau prover with proof reconstruction |
|
7 |
||
2854 | 8 |
SKOLEMIZES ReplaceI WRONGLY: allow new vars in prems, or forbid such rules?? |
2894 | 9 |
Needs explicit instantiation of assumptions? |
10 |
||
11 |
||
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
12 |
Blast_tac is often more powerful than fast_tac, but has some limitations. |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
13 |
Blast_tac... |
2894 | 14 |
* ignores addss, addbefore, addafter; this restriction is intrinsic |
15 |
* ignores elimination rules that don't have the correct format |
|
16 |
(conclusion must be a formula variable) |
|
17 |
* ignores types, which can make HOL proofs fail |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
18 |
* rules must not require higher-order unification, e.g. apply_type in ZF |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
19 |
+ message "Function Var's argument not a bound variable" relates to this |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
20 |
* its proof strategy is more general but can actually be slower |
2894 | 21 |
|
22 |
Known problems: |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
23 |
With substition for equalities (hyp_subst_tac): |
2894 | 24 |
1. Sometimes hyp_subst_tac will fail due to occurrence of the parameter |
25 |
as the argument of a function variable |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
26 |
2. When substitution affects a haz formula or literal, it is moved |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
27 |
back to the list of safe formulae. |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
28 |
But there's no way of putting it in the right place. A "moved" or |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
29 |
"no DETERM" flag would prevent proofs failing here. |
2854 | 30 |
*) |
31 |
||
32 |
||
33 |
(*Should be a type abbreviation?*) |
|
34 |
type netpair = (int*(bool*thm)) Net.net * (int*(bool*thm)) Net.net; |
|
35 |
||
36 |
||
37 |
(*Assumptions about constants: |
|
38 |
--The negation symbol is "Not" |
|
39 |
--The equality symbol is "op =" |
|
40 |
--The is-true judgement symbol is "Trueprop" |
|
41 |
--There are no constants named "*Goal* or "*False*" |
|
42 |
*) |
|
43 |
signature BLAST_DATA = |
|
44 |
sig |
|
45 |
type claset |
|
46 |
val notE : thm (* [| ~P; P |] ==> R *) |
|
47 |
val ccontr : thm |
|
48 |
val contr_tac : int -> tactic |
|
49 |
val dup_intr : thm -> thm |
|
50 |
val vars_gen_hyp_subst_tac : bool -> int -> tactic |
|
51 |
val claset : claset ref |
|
52 |
val rep_claset : |
|
53 |
claset -> {safeIs: thm list, safeEs: thm list, |
|
54 |
hazIs: thm list, hazEs: thm list, |
|
55 |
uwrapper: (int -> tactic) -> (int -> tactic), |
|
56 |
swrapper: (int -> tactic) -> (int -> tactic), |
|
57 |
safe0_netpair: netpair, safep_netpair: netpair, |
|
58 |
haz_netpair: netpair, dup_netpair: netpair} |
|
59 |
end; |
|
60 |
||
61 |
||
62 |
signature BLAST = |
|
63 |
sig |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
64 |
type claset |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
65 |
datatype term = |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
66 |
Const of string |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
67 |
| OConst of string * int |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
68 |
| Skolem of string * term option ref list |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
69 |
| Free of string |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
70 |
| Var of term option ref |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
71 |
| Bound of int |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
72 |
| Abs of string*term |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
73 |
| op $ of term*term; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
74 |
type branch |
2883 | 75 |
val depth_tac : claset -> int -> int -> tactic |
76 |
val blast_tac : claset -> int -> tactic |
|
77 |
val Blast_tac : int -> tactic |
|
78 |
val declConsts : string list * thm list -> unit |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
79 |
(*debugging tools*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
80 |
val trace : bool ref |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
81 |
val fullTrace : branch list list ref |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
82 |
val fromTerm : Type.type_sig -> Term.term -> term |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
83 |
val fromSubgoal : Type.type_sig -> Term.term -> term |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
84 |
val toTerm : int -> term -> Term.term |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
85 |
val readGoal : Sign.sg -> string -> term |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
86 |
val tryInThy : theory -> int -> string -> |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
87 |
branch list list * (int*int*exn) list * (int -> tactic) list |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
88 |
val trygl : claset -> int -> int -> |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
89 |
branch list list * (int*int*exn) list * (int -> tactic) list |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
90 |
val Trygl : int -> int -> |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
91 |
branch list list * (int*int*exn) list * (int -> tactic) list |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
92 |
val normBr : branch -> branch |
2854 | 93 |
end; |
94 |
||
95 |
||
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
96 |
functor BlastFun(Data: BLAST_DATA): BLAST = |
2854 | 97 |
struct |
98 |
||
99 |
type claset = Data.claset; |
|
100 |
||
101 |
val trace = ref false; |
|
102 |
||
103 |
datatype term = |
|
104 |
Const of string |
|
105 |
| OConst of string * int |
|
106 |
| Skolem of string * term option ref list |
|
107 |
| Free of string |
|
108 |
| Var of term option ref |
|
109 |
| Bound of int |
|
110 |
| Abs of string*term |
|
111 |
| op $ of term*term; |
|
112 |
||
113 |
||
114 |
exception DEST_EQ; |
|
115 |
||
116 |
(*Take apart an equality (plain or overloaded). NO constant Trueprop*) |
|
117 |
fun dest_eq (Const "op =" $ t $ u) = (t,u) |
|
118 |
| dest_eq (OConst("op =",_) $ t $ u) = (t,u) |
|
119 |
| dest_eq _ = raise DEST_EQ; |
|
120 |
||
121 |
(** Basic syntactic operations **) |
|
122 |
||
123 |
fun is_Var (Var _) = true |
|
124 |
| is_Var _ = false; |
|
125 |
||
126 |
fun dest_Var (Var x) = x; |
|
127 |
||
128 |
||
129 |
fun rand (f$x) = x; |
|
130 |
||
131 |
(* maps (f, [t1,...,tn]) to f(t1,...,tn) *) |
|
132 |
val list_comb : term * term list -> term = foldl (op $); |
|
133 |
||
134 |
||
135 |
(* maps f(t1,...,tn) to (f, [t1,...,tn]) ; naturally tail-recursive*) |
|
136 |
fun strip_comb u : term * term list = |
|
137 |
let fun stripc (f$t, ts) = stripc (f, t::ts) |
|
138 |
| stripc x = x |
|
139 |
in stripc(u,[]) end; |
|
140 |
||
141 |
||
142 |
(* maps f(t1,...,tn) to f , which is never a combination *) |
|
143 |
fun head_of (f$t) = head_of f |
|
144 |
| head_of u = u; |
|
145 |
||
146 |
||
147 |
(** Particular constants **) |
|
148 |
||
149 |
fun negate P = Const"Not" $ P; |
|
150 |
||
151 |
fun mkGoal P = Const"*Goal*" $ P; |
|
152 |
||
153 |
fun isGoal (Const"*Goal*" $ _) = true |
|
154 |
| isGoal _ = false; |
|
155 |
||
156 |
val Trueprop = Term.Const("Trueprop", Type("o",[])-->propT); |
|
157 |
fun mk_tprop P = Term.$ (Trueprop, P); |
|
158 |
||
159 |
fun isTrueprop (Term.Const("Trueprop",_)) = true |
|
160 |
| isTrueprop _ = false; |
|
161 |
||
162 |
||
163 |
(** Dealing with overloaded constants **) |
|
164 |
||
165 |
(*Result is a symbol table, indexed by names of overloaded constants. |
|
166 |
Each constant maps to a list of (pattern,Blast.Const) pairs. |
|
167 |
Any Term.Const that matches a pattern gets replaced by the Blast.Const. |
|
168 |
*) |
|
169 |
fun addConsts (t as Term.Const(a,_), tab) = |
|
170 |
(case Symtab.lookup (tab,a) of |
|
171 |
None => tab (*ignore: not a constant that we are looking for*) |
|
172 |
| Some patList => |
|
173 |
(case gen_assoc (op aconv) (patList, t) of |
|
174 |
None => Symtab.update |
|
175 |
((a, (t, OConst (a, length patList)) :: patList), |
|
176 |
tab) |
|
177 |
| _ => tab)) |
|
178 |
| addConsts (Term.Abs(_,_,body), tab) = addConsts (body, tab) |
|
179 |
| addConsts (Term.$ (t,u), tab) = addConsts (t, addConsts (u, tab)) |
|
180 |
| addConsts (_, tab) = tab (*ignore others*); |
|
181 |
||
182 |
||
183 |
fun addRules (rls,tab) = foldr addConsts (map (#prop o rep_thm) rls, tab); |
|
184 |
||
2883 | 185 |
fun declConst (a,tab) = |
186 |
case Symtab.lookup (tab,a) of |
|
187 |
None => Symtab.update((a,[]), tab) (*create a brand new entry*) |
|
188 |
| Some _ => tab (*preserve old entry*); |
|
2854 | 189 |
|
190 |
(*maps the name of each overloaded constant to a list of archetypal constants, |
|
191 |
which may be polymorphic.*) |
|
192 |
local |
|
193 |
val overLoadTab = ref (Symtab.null : (Term.term * term) list Symtab.table) |
|
194 |
(*The alists in this table should only be increased*) |
|
195 |
in |
|
196 |
||
197 |
fun declConsts (names, rls) = |
|
198 |
overLoadTab := addRules (rls, foldr declConst (names, !overLoadTab)); |
|
199 |
||
200 |
||
201 |
(*Convert a possibly overloaded Term.Const to a Blast.Const*) |
|
202 |
fun fromConst tsig (t as Term.Const (a,_)) = |
|
203 |
let fun find [] = Const a |
|
204 |
| find ((pat,t')::patList) = |
|
205 |
if Pattern.matches tsig (pat,t) then t' |
|
206 |
else find patList |
|
207 |
in case Symtab.lookup(!overLoadTab, a) of |
|
208 |
None => Const a |
|
209 |
| Some patList => find patList |
|
210 |
end; |
|
211 |
end; |
|
212 |
||
213 |
||
214 |
(*Tests whether 2 terms are alpha-convertible; chases instantiations*) |
|
215 |
fun (Const a) aconv (Const b) = a=b |
|
216 |
| (OConst ai) aconv (OConst bj) = ai=bj |
|
217 |
| (Skolem (a,_)) aconv (Skolem (b,_)) = a=b (*arglists must then be equal*) |
|
218 |
| (Free a) aconv (Free b) = a=b |
|
219 |
| (Var(ref(Some t))) aconv u = t aconv u |
|
220 |
| t aconv (Var(ref(Some u))) = t aconv u |
|
221 |
| (Var v) aconv (Var w) = v=w (*both Vars are un-assigned*) |
|
222 |
| (Bound i) aconv (Bound j) = i=j |
|
223 |
| (Abs(_,t)) aconv (Abs(_,u)) = t aconv u |
|
224 |
| (f$t) aconv (g$u) = (f aconv g) andalso (t aconv u) |
|
225 |
| _ aconv _ = false; |
|
226 |
||
227 |
||
228 |
fun mem_term (_, []) = false |
|
229 |
| mem_term (t, t'::ts) = t aconv t' orelse mem_term(t,ts); |
|
230 |
||
231 |
fun ins_term(t,ts) = if mem_term(t,ts) then ts else t :: ts; |
|
232 |
||
233 |
fun mem_var (v: term option ref, []) = false |
|
234 |
| mem_var (v, v'::vs) = v=v' orelse mem_var(v,vs); |
|
235 |
||
236 |
fun ins_var(v,vs) = if mem_var(v,vs) then vs else v :: vs; |
|
237 |
||
238 |
||
239 |
(** Vars **) |
|
240 |
||
241 |
(*Accumulates the Vars in the term, suppressing duplicates*) |
|
242 |
fun add_term_vars (Skolem(a,args), vars) = add_vars_vars(args,vars) |
|
243 |
| add_term_vars (Var (v as ref None), vars) = ins_var (v, vars) |
|
244 |
| add_term_vars (Var (ref (Some u)), vars) = add_term_vars(u,vars) |
|
245 |
| add_term_vars (Abs (_,body), vars) = add_term_vars(body,vars) |
|
246 |
| add_term_vars (f$t, vars) = add_term_vars (f, add_term_vars(t, vars)) |
|
247 |
| add_term_vars (_, vars) = vars |
|
248 |
(*Term list version. [The fold functionals are slow]*) |
|
249 |
and add_terms_vars ([], vars) = vars |
|
250 |
| add_terms_vars (t::ts, vars) = add_terms_vars (ts, add_term_vars(t,vars)) |
|
251 |
(*Var list version.*) |
|
252 |
and add_vars_vars ([], vars) = vars |
|
253 |
| add_vars_vars (ref (Some u) :: vs, vars) = |
|
254 |
add_vars_vars (vs, add_term_vars(u,vars)) |
|
255 |
| add_vars_vars (v::vs, vars) = (*v must be a ref None*) |
|
256 |
add_vars_vars (vs, ins_var (v, vars)); |
|
257 |
||
258 |
||
259 |
(*Chase assignments in "vars"; return a list of unassigned variables*) |
|
260 |
fun vars_in_vars vars = add_vars_vars(vars,[]); |
|
261 |
||
262 |
||
263 |
||
264 |
(*increment a term's non-local bound variables |
|
265 |
inc is increment for bound variables |
|
266 |
lev is level at which a bound variable is considered 'loose'*) |
|
267 |
fun incr_bv (inc, lev, u as Bound i) = if i>=lev then Bound(i+inc) else u |
|
268 |
| incr_bv (inc, lev, Abs(a,body)) = Abs(a, incr_bv(inc,lev+1,body)) |
|
269 |
| incr_bv (inc, lev, f$t) = incr_bv(inc,lev,f) $ incr_bv(inc,lev,t) |
|
270 |
| incr_bv (inc, lev, u) = u; |
|
271 |
||
272 |
fun incr_boundvars 0 t = t |
|
273 |
| incr_boundvars inc t = incr_bv(inc,0,t); |
|
274 |
||
275 |
||
276 |
(*Accumulate all 'loose' bound vars referring to level 'lev' or beyond. |
|
277 |
(Bound 0) is loose at level 0 *) |
|
278 |
fun add_loose_bnos (Bound i, lev, js) = if i<lev then js |
|
279 |
else (i-lev) ins_int js |
|
280 |
| add_loose_bnos (Abs (_,t), lev, js) = add_loose_bnos (t, lev+1, js) |
|
281 |
| add_loose_bnos (f$t, lev, js) = |
|
282 |
add_loose_bnos (f, lev, add_loose_bnos (t, lev, js)) |
|
283 |
| add_loose_bnos (_, _, js) = js; |
|
284 |
||
285 |
fun loose_bnos t = add_loose_bnos (t, 0, []); |
|
286 |
||
287 |
fun subst_bound (arg, t) : term = |
|
288 |
let fun subst (t as Bound i, lev) = |
|
289 |
if i<lev then t (*var is locally bound*) |
|
290 |
else if i=lev then incr_boundvars lev arg |
|
291 |
else Bound(i-1) (*loose: change it*) |
|
292 |
| subst (Abs(a,body), lev) = Abs(a, subst(body,lev+1)) |
|
293 |
| subst (f$t, lev) = subst(f,lev) $ subst(t,lev) |
|
294 |
| subst (t,lev) = t |
|
295 |
in subst (t,0) end; |
|
296 |
||
297 |
||
298 |
(*Normalize...but not the bodies of ABSTRACTIONS*) |
|
299 |
fun norm t = case t of |
|
300 |
Skolem(a,args) => Skolem(a, vars_in_vars args) |
|
301 |
| (Var (ref None)) => t |
|
302 |
| (Var (ref (Some u))) => norm u |
|
303 |
| (f $ u) => (case norm f of |
|
304 |
Abs(_,body) => norm (subst_bound (u, body)) |
|
305 |
| nf => nf $ norm u) |
|
306 |
| _ => t; |
|
307 |
||
308 |
||
309 |
(*Weak (one-level) normalize for use in unification*) |
|
310 |
fun wkNormAux t = case t of |
|
311 |
(Var v) => (case !v of |
|
312 |
Some u => wkNorm u |
|
313 |
| None => t) |
|
314 |
| (f $ u) => (case wkNormAux f of |
|
315 |
Abs(_,body) => wkNorm (subst_bound (u, body)) |
|
316 |
| nf => nf $ u) |
|
317 |
| _ => t |
|
318 |
and wkNorm t = case head_of t of |
|
319 |
Const _ => t |
|
320 |
| OConst _ => t |
|
321 |
| Skolem(a,args) => t |
|
322 |
| Free _ => t |
|
323 |
| _ => wkNormAux t; |
|
324 |
||
325 |
||
326 |
(*Does variable v occur in u? For unification.*) |
|
327 |
fun varOccur v = |
|
328 |
let fun occL [] = false (*same as (exists occ), but faster*) |
|
329 |
| occL (u::us) = occ u orelse occL us |
|
330 |
and occ (Var w) = |
|
331 |
v=w orelse |
|
332 |
(case !w of None => false |
|
333 |
| Some u => occ u) |
|
334 |
| occ (Skolem(_,args)) = occL (map Var args) |
|
335 |
| occ (Abs(_,u)) = occ u |
|
336 |
| occ (f$u) = occ u orelse occ f |
|
337 |
| occ (_) = false; |
|
338 |
in occ end; |
|
339 |
||
340 |
exception UNIFY; |
|
341 |
||
342 |
val trail = ref [] : term option ref list ref; |
|
343 |
val ntrail = ref 0; |
|
344 |
||
345 |
||
346 |
(*Restore the trail to some previous state: for backtracking*) |
|
347 |
fun clearTo n = |
|
348 |
while !ntrail>n do |
|
349 |
(hd(!trail) := None; |
|
350 |
trail := tl (!trail); |
|
351 |
ntrail := !ntrail - 1); |
|
352 |
||
353 |
||
354 |
(*First-order unification with bound variables. |
|
355 |
"vars" is a list of variables local to the rule and NOT to be put |
|
356 |
on the trail (no point in doing so) |
|
357 |
*) |
|
358 |
fun unify(vars,t,u) = |
|
359 |
let val n = !ntrail |
|
360 |
fun update (t as Var v, u) = |
|
361 |
if t aconv u then () |
|
362 |
else if varOccur v u then raise UNIFY |
|
363 |
else if mem_var(v, vars) then v := Some u |
|
364 |
else (*avoid updating Vars in the branch if possible!*) |
|
365 |
if is_Var u andalso mem_var(dest_Var u, vars) |
|
366 |
then dest_Var u := Some t |
|
367 |
else (v := Some u; |
|
368 |
trail := v :: !trail; ntrail := !ntrail + 1) |
|
369 |
fun unifyAux (t,u) = |
|
370 |
case (wkNorm t, wkNorm u) of |
|
371 |
(nt as Var v, nu) => update(nt,nu) |
|
372 |
| (nu, nt as Var v) => update(nt,nu) |
|
373 |
| (Abs(_,t'), Abs(_,u')) => unifyAux(t',u') |
|
374 |
(*NB: can yield unifiers having dangling Bound vars!*) |
|
375 |
| (f$t', g$u') => (unifyAux(f,g); unifyAux(t',u')) |
|
376 |
| (nt, nu) => if nt aconv nu then () else raise UNIFY |
|
377 |
in unifyAux(t,u) handle UNIFY => (clearTo n; raise UNIFY) |
|
378 |
end; |
|
379 |
||
380 |
||
381 |
(*Convert from "real" terms to prototerms; eta-contract*) |
|
382 |
fun fromTerm tsig t = |
|
383 |
let val alist = ref [] |
|
384 |
fun from (t as Term.Const _) = fromConst tsig t |
|
385 |
| from (Term.Free (a,_)) = Free a |
|
386 |
| from (Term.Bound i) = Bound i |
|
387 |
| from (Term.Var (ixn,T)) = |
|
388 |
(case (assoc_string_int(!alist,ixn)) of |
|
389 |
None => let val t' = Var(ref None) |
|
390 |
in alist := (ixn, (t', T)) :: !alist; t' |
|
391 |
end |
|
392 |
| Some (v,_) => v) |
|
393 |
| from (Term.Abs (a,_,u)) = |
|
394 |
(case from u of |
|
395 |
u' as (f $ Bound 0) => |
|
396 |
if (0 mem_int loose_bnos f) then Abs(a,u') |
|
397 |
else incr_boundvars ~1 f |
|
398 |
| u' => Abs(a,u')) |
|
399 |
| from (Term.$ (f,u)) = from f $ from u |
|
400 |
in from t end; |
|
401 |
||
402 |
(* A1==>...An==>B goes to [A1,...,An], where B is not an implication *) |
|
403 |
fun strip_imp_prems (Const"==>" $ (Const"Trueprop" $ A) $ B) = |
|
404 |
A :: strip_imp_prems B |
|
405 |
| strip_imp_prems (Const"==>" $ A $ B) = A :: strip_imp_prems B |
|
406 |
| strip_imp_prems _ = []; |
|
407 |
||
408 |
(* A1==>...An==>B goes to B, where B is not an implication *) |
|
409 |
fun strip_imp_concl (Const"==>" $ A $ B) = strip_imp_concl B |
|
410 |
| strip_imp_concl (Const"Trueprop" $ A) = A |
|
411 |
| strip_imp_concl A = A : term; |
|
412 |
||
413 |
||
414 |
(*** Conversion of Elimination Rules to Tableau Operations ***) |
|
415 |
||
416 |
(*The conclusion becomes the goal/negated assumption *False*: delete it!*) |
|
417 |
fun squash_nots [] = [] |
|
418 |
| squash_nots (Const "*Goal*" $ (Var (ref (Some (Const"*False*")))) :: Ps) = |
|
419 |
squash_nots Ps |
|
420 |
| squash_nots (Const "Not" $ (Var (ref (Some (Const"*False*")))) :: Ps) = |
|
421 |
squash_nots Ps |
|
422 |
| squash_nots (P::Ps) = P :: squash_nots Ps; |
|
423 |
||
424 |
fun skoPrem vars (Const "all" $ Abs (_, P)) = |
|
425 |
skoPrem vars (subst_bound (Skolem (gensym "S_", vars), P)) |
|
426 |
| skoPrem vars P = P; |
|
427 |
||
428 |
fun convertPrem t = |
|
429 |
squash_nots (mkGoal (strip_imp_concl t) :: strip_imp_prems t); |
|
430 |
||
431 |
(*Expects elimination rules to have a formula variable as conclusion*) |
|
432 |
fun convertRule vars t = |
|
433 |
let val (P::Ps) = strip_imp_prems t |
|
434 |
val Var v = strip_imp_concl t |
|
435 |
in v := Some (Const"*False*"); |
|
436 |
(P, map (convertPrem o skoPrem vars) Ps) |
|
437 |
end; |
|
438 |
||
439 |
||
440 |
(*Like dup_elim, but puts the duplicated major premise FIRST*) |
|
441 |
fun rev_dup_elim th = th RSN (2, revcut_rl) |> assumption 2 |> Sequence.hd; |
|
442 |
||
443 |
||
444 |
(*Count new hyps so that they can be rotated*) |
|
445 |
fun nNewHyps [] = 0 |
|
446 |
| nNewHyps (Const "*Goal*" $ _ :: Ps) = nNewHyps Ps |
|
447 |
| nNewHyps (P::Ps) = 1 + nNewHyps Ps; |
|
448 |
||
449 |
fun rot_subgoals_tac [] i st = Sequence.single st |
|
450 |
| rot_subgoals_tac (k::ks) i st = |
|
451 |
rot_subgoals_tac ks (i+1) (Sequence.hd (rotate_tac (~k) i st)) |
|
452 |
handle OPTION _ => Sequence.null; |
|
453 |
||
454 |
fun TRACE rl tac st = if !trace then (prth rl; tac st) else tac st; |
|
455 |
||
456 |
(*Tableau rule from elimination rule. Flag "dup" requests duplication of the |
|
457 |
affected formula.*) |
|
458 |
fun fromRule vars rl = |
|
459 |
let val {tsig,...} = Sign.rep_sg (#sign (rep_thm rl)) |
|
460 |
val trl = rl |> rep_thm |> #prop |> fromTerm tsig |> convertRule vars |
|
461 |
fun tac dup i = |
|
462 |
TRACE rl |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
463 |
(etac (if dup then rev_dup_elim rl else rl) i) |
2854 | 464 |
THEN rot_subgoals_tac (map nNewHyps (#2 trl)) i |
465 |
||
466 |
in General.SOME (trl, tac) end |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
467 |
handle Bind => (*reject: conclusion is not just a variable*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
468 |
(if !trace then (writeln"Warning: ignoring ill-formed elimination rule"; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
469 |
print_thm rl) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
470 |
else (); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
471 |
General.NONE); |
2854 | 472 |
|
473 |
||
474 |
(*** Conversion of Introduction Rules (needed for efficiency in |
|
475 |
proof reconstruction) ***) |
|
476 |
||
477 |
fun convertIntrPrem t = mkGoal (strip_imp_concl t) :: strip_imp_prems t; |
|
478 |
||
479 |
fun convertIntrRule vars t = |
|
480 |
let val Ps = strip_imp_prems t |
|
481 |
val P = strip_imp_concl t |
|
482 |
in (mkGoal P, map (convertIntrPrem o skoPrem vars) Ps) |
|
483 |
end; |
|
484 |
||
485 |
(*Tableau rule from introduction rule. Since haz rules are now delayed, |
|
486 |
"dup" is always FALSE for introduction rules.*) |
|
487 |
fun fromIntrRule vars rl = |
|
488 |
let val {tsig,...} = Sign.rep_sg (#sign (rep_thm rl)) |
|
489 |
val trl = rl |> rep_thm |> #prop |> fromTerm tsig |> convertIntrRule vars |
|
490 |
fun tac dup i = |
|
491 |
TRACE rl (DETERM (rtac (if dup then Data.dup_intr rl else rl) i)) |
|
492 |
THEN rot_subgoals_tac (map nNewHyps (#2 trl)) i |
|
493 |
in (trl, tac) end; |
|
494 |
||
495 |
||
496 |
val dummyVar = Term.Var (("Doom",666), dummyT); |
|
497 |
||
498 |
(*Convert from prototerms to ordinary terms with dummy types |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
499 |
Ignore abstractions; identify all Vars; STOP at given depth*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
500 |
fun toTerm 0 _ = dummyVar |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
501 |
| toTerm d (Const a) = Term.Const (a,dummyT) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
502 |
| toTerm d (OConst(a,_)) = Term.Const (a,dummyT) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
503 |
| toTerm d (Skolem(a,_)) = Term.Const (a,dummyT) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
504 |
| toTerm d (Free a) = Term.Free (a,dummyT) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
505 |
| toTerm d (Bound i) = Term.Bound i |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
506 |
| toTerm d (Var _) = dummyVar |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
507 |
| toTerm d (Abs(a,_)) = dummyVar |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
508 |
| toTerm d (f $ u) = Term.$ (toTerm d f, toTerm (d-1) u); |
2854 | 509 |
|
510 |
||
511 |
fun netMkRules P vars (nps: netpair list) = |
|
512 |
case P of |
|
513 |
(Const "*Goal*" $ G) => |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
514 |
let val pG = mk_tprop (toTerm 2 G) |
2854 | 515 |
val intrs = List.concat |
516 |
(map (fn (inet,_) => Net.unify_term inet pG) |
|
517 |
nps) |
|
518 |
in map (fromIntrRule vars o #2) (orderlist intrs) end |
|
519 |
| _ => |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
520 |
let val pP = mk_tprop (toTerm 3 P) |
2854 | 521 |
val elims = List.concat |
522 |
(map (fn (_,enet) => Net.unify_term enet pP) |
|
523 |
nps) |
|
524 |
in List.mapPartial (fromRule vars o #2) (orderlist elims) end; |
|
525 |
||
526 |
(** |
|
527 |
end; |
|
528 |
**) |
|
529 |
||
530 |
(*** Code for handling equality: naive substitution, like hyp_subst_tac ***) |
|
531 |
||
532 |
(*Replace the ATOMIC term "old" by "new" in t*) |
|
533 |
fun subst_atomic (old,new) t = |
|
534 |
let fun subst (Var(ref(Some u))) = subst u |
|
535 |
| subst (Abs(a,body)) = Abs(a, subst body) |
|
536 |
| subst (f$t) = subst f $ subst t |
|
537 |
| subst t = if t aconv old then new else t |
|
538 |
in subst t end; |
|
539 |
||
540 |
(*Eta-contract a term from outside: just enough to reduce it to an atom*) |
|
541 |
fun eta_contract_atom (t0 as Abs(a, body)) = |
|
542 |
(case eta_contract2 body of |
|
543 |
f $ Bound 0 => if (0 mem_int loose_bnos f) then t0 |
|
544 |
else eta_contract_atom (incr_boundvars ~1 f) |
|
545 |
| _ => t0) |
|
546 |
| eta_contract_atom t = t |
|
547 |
and eta_contract2 (f$t) = f $ eta_contract_atom t |
|
548 |
| eta_contract2 t = eta_contract_atom t; |
|
549 |
||
550 |
||
551 |
(*When can we safely delete the equality? |
|
552 |
Not if it equates two constants; consider 0=1. |
|
553 |
Not if it resembles x=t[x], since substitution does not eliminate x. |
|
554 |
Not if it resembles ?x=0; another goal could instantiate ?x to Suc(i) |
|
555 |
Prefer to eliminate Bound variables if possible. |
|
556 |
Result: true = use as is, false = reorient first *) |
|
557 |
||
558 |
(*Does t occur in u? For substitution. |
|
559 |
Does NOT check args of Skolem terms: substitution does not affect them. |
|
560 |
NOT reflexive since hyp_subst_tac fails on x=x.*) |
|
561 |
fun substOccur t = |
|
562 |
let fun occEq u = (t aconv u) orelse occ u |
|
563 |
and occ (Var(ref None)) = false |
|
564 |
| occ (Var(ref(Some u))) = occEq u |
|
565 |
| occ (Abs(_,u)) = occEq u |
|
566 |
| occ (f$u) = occEq u orelse occEq f |
|
567 |
| occ (_) = false; |
|
568 |
in occEq end; |
|
569 |
||
570 |
fun check (t,u,v) = if substOccur t u then raise DEST_EQ else v; |
|
571 |
||
572 |
(*IF the goal is an equality with a substitutable variable |
|
573 |
THEN orient that equality ELSE raise exception DEST_EQ*) |
|
574 |
fun orientGoal (t,u) = |
|
575 |
case (eta_contract_atom t, eta_contract_atom u) of |
|
576 |
(Skolem _, _) => check(t,u,(t,u)) (*eliminates t*) |
|
577 |
| (_, Skolem _) => check(u,t,(u,t)) (*eliminates u*) |
|
578 |
| (Free _, _) => check(t,u,(t,u)) (*eliminates t*) |
|
579 |
| (_, Free _) => check(u,t,(u,t)) (*eliminates u*) |
|
580 |
| _ => raise DEST_EQ; |
|
581 |
||
582 |
||
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
583 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
584 |
|
2894 | 585 |
(*Substitute through the branch if an equality goal (else raise DEST_EQ). |
586 |
Moves affected literals back into the branch, but it is not clear where |
|
587 |
they should go: this could make proofs fail. ??? *) |
|
588 |
fun equalSubst (G, pairs, lits, vars, lim) = |
|
2854 | 589 |
let val subst = subst_atomic (orientGoal(dest_eq G)) |
590 |
fun subst2(G,md) = (subst G, md) |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
591 |
(*substitute throughout Haz list; move affected ones to Safe part*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
592 |
fun subHazs ([], Gs, nHs) = (Gs,nHs) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
593 |
| subHazs ((H,md)::Hs, Gs, nHs) = |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
594 |
let val nH = subst H |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
595 |
in if nH aconv H then subHazs (Hs, Gs, (H,md)::nHs) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
596 |
else subHazs (Hs, (nH,md)::Gs, nHs) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
597 |
end |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
598 |
val pairs' = map (fn(Gs,Hs) => subHazs(rev Hs, map subst2 Gs, [])) pairs |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
599 |
(*substitute throughout literals; move affected ones to Safe part*) |
2894 | 600 |
fun subLits ([], [], nlits) = (pairs', nlits, vars, lim) |
601 |
| subLits ([], Gs, nlits) = ((Gs,[])::pairs', nlits, vars, lim) |
|
602 |
| subLits (lit::lits, Gs, nlits) = |
|
2854 | 603 |
let val nlit = subst lit |
2894 | 604 |
in if nlit aconv lit then subLits (lits, Gs, nlit::nlits) |
605 |
else subLits (lits, (nlit,true)::Gs, nlits) |
|
2854 | 606 |
end |
2894 | 607 |
in subLits (rev lits, [], []) |
2854 | 608 |
end; |
609 |
||
610 |
||
611 |
exception NEWBRANCHES and CLOSEF; |
|
612 |
||
2894 | 613 |
(*Pending formulae carry md (may duplicate) flags*) |
614 |
type branch = ((term*bool) list * (*safe formulae on this level*) |
|
615 |
(term*bool) list) list * (*haz formulae on this level*) |
|
2854 | 616 |
term list * (*literals: irreducible formulae*) |
617 |
term option ref list * (*variables occurring in branch*) |
|
618 |
int; (*resource limit*) |
|
619 |
||
620 |
val fullTrace = ref[] : branch list list ref; |
|
621 |
||
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
622 |
(*Normalize a branch--for tracing*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
623 |
local |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
624 |
fun norm2 (G,md) = (norm G, md); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
625 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
626 |
fun normLev (Gs,Hs) = (map norm2 Gs, map norm2 Hs); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
627 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
628 |
in |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
629 |
fun normBr (br, lits, vars, lim) = |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
630 |
(map normLev br, map norm lits, vars, lim); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
631 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
632 |
fun tracing brs = if !trace then fullTrace := map normBr brs :: !fullTrace |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
633 |
else () |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
634 |
end; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
635 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
636 |
|
2854 | 637 |
exception PROVE; |
638 |
||
639 |
val eq_contr_tac = eresolve_tac [Data.notE] THEN' eq_assume_tac; |
|
640 |
||
641 |
val eContr_tac = TRACE Data.notE (eq_contr_tac ORELSE' Data.contr_tac); |
|
642 |
val eAssume_tac = TRACE asm_rl (eq_assume_tac ORELSE' assume_tac); |
|
643 |
||
644 |
(*Try to unify complementary literals and return the corresponding tactic. *) |
|
645 |
fun tryClose (Const"*Goal*" $ G, L) = (unify([],G,L); eAssume_tac) |
|
646 |
| tryClose (G, Const"*Goal*" $ L) = (unify([],G,L); eAssume_tac) |
|
647 |
| tryClose (Const"Not" $ G, L) = (unify([],G,L); eContr_tac) |
|
648 |
| tryClose (G, Const"Not" $ L) = (unify([],G,L); eContr_tac) |
|
649 |
| tryClose _ = raise UNIFY; |
|
650 |
||
651 |
||
652 |
(*Were there Skolem terms in the premise? Must NOT chase Vars*) |
|
653 |
fun hasSkolem (Skolem _) = true |
|
654 |
| hasSkolem (Abs (_,body)) = hasSkolem body |
|
655 |
| hasSkolem (f$t) = hasSkolem f orelse hasSkolem t |
|
656 |
| hasSkolem _ = false; |
|
657 |
||
658 |
(*Attach the right "may duplicate" flag to new formulae: if they contain |
|
659 |
Skolem terms then allow duplication.*) |
|
660 |
fun joinMd md [] = [] |
|
661 |
| joinMd md (G::Gs) = (G, hasSkolem G orelse md) :: joinMd md Gs; |
|
662 |
||
2894 | 663 |
(*Convert a Goal to an ordinary Not. Used also in dup_intr, where a goal like |
664 |
Ex(P) is duplicated as the assumption ~Ex(P). *) |
|
665 |
fun negOfGoal (Const"*Goal*" $ G) = negate G |
|
666 |
| negOfGoal G = G; |
|
667 |
||
668 |
fun negOfGoal2 (G,md) = (negOfGoal G, md); |
|
669 |
||
670 |
(*Converts all Goals to Nots in the safe parts of a branch. They could |
|
671 |
have been moved there from the literals list after substitution (equalSubst). |
|
672 |
There can be at most one--this function could be made more efficient.*) |
|
673 |
fun negOfGoals pairs = map (fn (Gs,haz) => (map negOfGoal2 Gs, haz)) pairs; |
|
674 |
||
675 |
(*Tactic. Convert *Goal* to negated assumption in FIRST position*) |
|
676 |
val negOfGoal_tac = rtac Data.ccontr THEN' rotate_tac ~1; |
|
677 |
||
2854 | 678 |
|
679 |
(** Backtracking and Pruning **) |
|
680 |
||
681 |
(*clashVar vars (n,trail) determines whether any of the last n elements |
|
682 |
of "trail" occur in "vars" OR in their instantiations*) |
|
683 |
fun clashVar [] = (fn _ => false) |
|
684 |
| clashVar vars = |
|
685 |
let fun clash (0, _) = false |
|
686 |
| clash (_, []) = false |
|
687 |
| clash (n, v::vs) = exists (varOccur v) vars orelse clash(n-1,vs) |
|
688 |
in clash end; |
|
689 |
||
690 |
||
691 |
(*nbrs = # of branches just prior to closing this one. Delete choice points |
|
692 |
for goals proved by the latest inference, provided NO variables in the |
|
693 |
next branch have been updated.*) |
|
694 |
fun prune (1, nxtVars, choices) = choices (*DON'T prune at very end: allow |
|
695 |
backtracking over bad proofs*) |
|
696 |
| prune (nbrs, nxtVars, choices) = |
|
697 |
let fun traceIt last = |
|
698 |
let val ll = length last |
|
699 |
and lc = length choices |
|
700 |
in if !trace andalso ll<lc then |
|
701 |
(writeln("PRUNING " ^ Int.toString(lc-ll) ^ " LEVELS"); |
|
702 |
last) |
|
703 |
else last |
|
704 |
end |
|
705 |
fun pruneAux (last, _, _, []) = last |
|
706 |
| pruneAux (last, ntrl, trl, ch' as (ntrl',nbrs',exn) :: choices) = |
|
707 |
if nbrs' < nbrs |
|
708 |
then last (*don't backtrack beyond first solution of goal*) |
|
709 |
else if nbrs' > nbrs then pruneAux (last, ntrl, trl, choices) |
|
710 |
else (* nbrs'=nbrs *) |
|
711 |
if clashVar nxtVars (ntrl-ntrl', trl) then last |
|
712 |
else (*no clashes: can go back at least this far!*) |
|
713 |
pruneAux(choices, ntrl', List.drop(trl, ntrl-ntrl'), |
|
714 |
choices) |
|
715 |
in traceIt (pruneAux (choices, !ntrail, !trail, choices)) end; |
|
716 |
||
2894 | 717 |
fun nextVars ((br, lits, vars, lim) :: _) = map Var vars |
2854 | 718 |
| nextVars [] = []; |
719 |
||
720 |
fun backtrack ((_, _, exn)::_) = raise exn |
|
721 |
| backtrack _ = raise PROVE; |
|
722 |
||
2894 | 723 |
(*Add the literal G, handling *Goal* and detecting duplicates.*) |
724 |
fun addLit (Const "*Goal*" $ G, lits) = |
|
725 |
(*New literal is a *Goal*, so change all other Goals to Nots*) |
|
2854 | 726 |
let fun bad (Const"*Goal*" $ _) = true |
727 |
| bad (Const"Not" $ G') = G aconv G' |
|
728 |
| bad _ = false; |
|
729 |
fun change [] = [] |
|
730 |
| change (Const"*Goal*" $ G' :: lits) = |
|
731 |
if G aconv G' then change lits |
|
732 |
else Const"Not" $ G' :: change lits |
|
733 |
| change (Const"Not" $ G' :: lits) = |
|
734 |
if G aconv G' then change lits |
|
735 |
else Const"Not" $ G' :: change lits |
|
736 |
| change (lit::lits) = lit :: change lits |
|
737 |
in |
|
738 |
Const "*Goal*" $ G :: (if exists bad lits then change lits else lits) |
|
739 |
end |
|
740 |
| addLit (G,lits) = ins_term(G, lits) |
|
741 |
||
742 |
||
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
743 |
(*For calculating the "penalty" to assess on a branching factor of n*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
744 |
fun log2 1 = 0 |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
745 |
| log2 n = 1 + log2(n div 2); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
746 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
747 |
|
2854 | 748 |
(*Tableau prover based on leanTaP. Argument is a list of branches. Each |
749 |
branch contains a list of unexpanded formulae, a list of literals, and a |
|
750 |
bound on unsafe expansions.*) |
|
751 |
fun prove (cs, brs, cont) = |
|
752 |
let val {safe0_netpair, safep_netpair, haz_netpair, ...} = Data.rep_claset cs |
|
753 |
val safeList = [safe0_netpair, safep_netpair] |
|
754 |
and hazList = [haz_netpair] |
|
755 |
fun prv (tacs, trs, choices, []) = (cont (trs,choices,tacs)) |
|
756 |
| prv (tacs, trs, choices, |
|
2894 | 757 |
brs0 as (((G,md)::br, haz)::pairs, lits, vars, lim) :: brs) = |
2854 | 758 |
let exception PRV (*backtrack to precisely this recursion!*) |
759 |
val ntrl = !ntrail |
|
760 |
val nbrs = length brs0 |
|
761 |
val nxtVars = nextVars brs |
|
762 |
val G = norm G |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
763 |
val rules = netMkRules G vars safeList |
2854 | 764 |
(*Make a new branch, decrementing "lim" if instantiations occur*) |
2894 | 765 |
fun newBr (vars',lim') prems = |
766 |
map (fn prem => |
|
767 |
if (exists isGoal prem) |
|
768 |
then (((joinMd md prem, []) :: |
|
769 |
negOfGoals ((br, haz)::pairs)), |
|
770 |
map negOfGoal lits, |
|
771 |
vars', lim') |
|
772 |
else (((joinMd md prem, []) :: (br, haz) :: pairs), |
|
773 |
lits, vars', lim')) |
|
2854 | 774 |
prems @ |
775 |
brs |
|
776 |
(*Seek a matching rule. If unifiable then add new premises |
|
777 |
to branch.*) |
|
778 |
fun deeper [] = raise NEWBRANCHES |
|
779 |
| deeper (((P,prems),tac)::grls) = |
|
780 |
let val dummy = unify(add_term_vars(P,[]), P, G) |
|
781 |
(*P comes from the rule; G comes from the branch.*) |
|
782 |
val ntrl' = !ntrail |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
783 |
val lim' = if ntrl < !ntrail |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
784 |
then lim - (1+log2(length rules)) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
785 |
else lim (*discourage branching updates*) |
2894 | 786 |
val vars = vars_in_vars vars |
787 |
val vars' = foldr add_terms_vars (prems, vars) |
|
2854 | 788 |
val choices' = (ntrl, nbrs, PRV) :: choices |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
789 |
val tacs' = (DETERM o (tac false)) :: tacs |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
790 |
(*no duplication*) |
2854 | 791 |
in |
792 |
if null prems then (*closed the branch: prune!*) |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
793 |
prv(tacs', brs0::trs, |
2854 | 794 |
prune (nbrs, nxtVars, choices'), |
795 |
brs) |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
796 |
handle PRV => (*reset Vars and try another rule*) |
2854 | 797 |
(clearTo ntrl; deeper grls) |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
798 |
else (*can we kill all the alternatives??*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
799 |
if lim'<0 then raise NEWBRANCHES |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
800 |
else |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
801 |
prv(tacs', brs0::trs, choices', |
2894 | 802 |
newBr (vars',lim') prems) |
2854 | 803 |
handle PRV => |
804 |
if ntrl < ntrl' then |
|
805 |
(*Vars have been updated: must backtrack |
|
806 |
even if not mentioned in other goals! |
|
807 |
Reset Vars and try another rule*) |
|
808 |
(clearTo ntrl; deeper grls) |
|
809 |
else (*backtrack to previous level*) |
|
810 |
backtrack choices |
|
811 |
end |
|
812 |
handle UNIFY => deeper grls |
|
813 |
(*Try to close branch by unifying with head goal*) |
|
814 |
fun closeF [] = raise CLOSEF |
|
815 |
| closeF (L::Ls) = |
|
816 |
let val tacs' = tryClose(G,L)::tacs |
|
817 |
val choices' = prune (nbrs, nxtVars, |
|
818 |
(ntrl, nbrs, PRV) :: choices) |
|
819 |
in prv (tacs', brs0::trs, choices', brs) |
|
820 |
handle PRV => |
|
821 |
(*reset Vars and try another literal |
|
822 |
[this handler is pruned if possible!]*) |
|
823 |
(clearTo ntrl; closeF Ls) |
|
824 |
end |
|
825 |
handle UNIFY => closeF Ls |
|
2894 | 826 |
fun closeFl [] = raise CLOSEF |
827 |
| closeFl ((br, haz)::pairs) = |
|
828 |
closeF (map fst br) |
|
829 |
handle CLOSEF => closeF (map fst haz) |
|
830 |
handle CLOSEF => closeFl pairs |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
831 |
in tracing brs0; |
2854 | 832 |
if lim<0 then backtrack choices |
833 |
else |
|
834 |
prv (Data.vars_gen_hyp_subst_tac false :: tacs, |
|
835 |
brs0::trs, choices, |
|
2894 | 836 |
equalSubst (G, (br,haz)::pairs, lits, vars, lim) :: brs) |
2854 | 837 |
handle DEST_EQ => closeF lits |
2894 | 838 |
handle CLOSEF => closeFl ((br,haz)::pairs) |
2854 | 839 |
handle CLOSEF => |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
840 |
deeper rules |
2894 | 841 |
handle NEWBRANCHES => |
842 |
(case netMkRules G vars hazList of |
|
843 |
[] => (*no plausible rules: move G to literals*) |
|
844 |
prv (tacs, brs0::trs, choices, |
|
845 |
((br,haz)::pairs, addLit(G,lits), vars, lim) |
|
846 |
::brs) |
|
847 |
| _ => (*G admits some haz rules: try later*) |
|
848 |
prv (if isGoal G then negOfGoal_tac :: tacs |
|
849 |
else tacs, |
|
850 |
brs0::trs, choices, |
|
851 |
((br, haz@[(negOfGoal G, md)])::pairs, |
|
852 |
lits, vars, lim) :: brs)) |
|
2854 | 853 |
end |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
854 |
| prv (tacs, trs, choices, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
855 |
(([],haz)::(Gs,haz')::pairs, lits, vars, lim) :: brs) = |
2894 | 856 |
(*no more "safe" formulae: transfer haz down a level*) |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
857 |
prv (tacs, trs, choices, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
858 |
((Gs,haz@haz')::pairs, lits, vars, lim) :: brs) |
2854 | 859 |
| prv (tacs, trs, choices, |
2894 | 860 |
brs0 as ([([], (H,md)::Hs)], lits, vars, lim) :: brs) = |
861 |
(*no safe steps possible at any level: apply a haz rule*) |
|
2854 | 862 |
let exception PRV (*backtrack to precisely this recursion!*) |
2894 | 863 |
val H = norm H |
2854 | 864 |
val ntrl = !ntrail |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
865 |
val rules = netMkRules H vars hazList |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
866 |
fun newPrem (vars,dup,lim') prem = |
2894 | 867 |
([(map (fn P => (P,false)) prem, []), (*may NOT duplicate*) |
868 |
([], if dup then Hs @ [(negOfGoal H, md)] else Hs)], |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
869 |
lits, vars, lim') |
2854 | 870 |
fun newBr x prems = map (newPrem x) prems @ brs |
871 |
(*Seek a matching rule. If unifiable then add new premises |
|
872 |
to branch.*) |
|
873 |
fun deeper [] = raise NEWBRANCHES |
|
874 |
| deeper (((P,prems),tac)::grls) = |
|
2894 | 875 |
let val dummy = unify(add_term_vars(P,[]), P, H) |
2854 | 876 |
val ntrl' = !ntrail |
877 |
val vars = vars_in_vars vars |
|
878 |
val vars' = foldr add_terms_vars (prems, vars) |
|
879 |
val dup = md andalso vars' <> vars |
|
2894 | 880 |
(*duplicate H only if md and the premise has new vars*) |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
881 |
val lim' = (*Decrement "lim" extra if updates occur*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
882 |
if ntrl < !ntrail then lim - (1+log2(length rules)) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
883 |
else lim-1 |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
884 |
(*NB: if the formula is duplicated, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
885 |
then it seems plausible to leave lim alone. |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
886 |
But then transitivity laws loop.*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
887 |
val mayUndo = not (null grls) (*other rules to try?*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
888 |
orelse ntrl < ntrl' (*variables updated?*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
889 |
orelse vars=vars' (*no new Vars?*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
890 |
val tac' = if mayUndo then tac dup |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
891 |
else DETERM o (tac dup) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
892 |
in |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
893 |
(*can we kill all the alternatives??*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
894 |
if lim'<0 andalso not (null prems) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
895 |
then raise NEWBRANCHES |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
896 |
else |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
897 |
prv(tac dup :: tacs, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
898 |
brs0::trs, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
899 |
(ntrl, length brs0, PRV) :: choices, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
900 |
newBr (vars', dup, lim') prems) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
901 |
handle PRV => |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
902 |
if mayUndo |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
903 |
then (*reset Vars and try another rule*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
904 |
(clearTo ntrl; deeper grls) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
905 |
else (*backtrack to previous level*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
906 |
backtrack choices |
2854 | 907 |
end |
908 |
handle UNIFY => deeper grls |
|
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
909 |
in tracing brs0; |
2854 | 910 |
if lim<1 then backtrack choices |
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
911 |
else deeper rules |
2854 | 912 |
handle NEWBRANCHES => |
2894 | 913 |
(*cannot close branch: move H to literals*) |
2854 | 914 |
prv (tacs, brs0::trs, choices, |
2894 | 915 |
([([], Hs)], H::lits, vars, lim)::brs) |
2854 | 916 |
end |
917 |
| prv (tacs, trs, choices, _ :: brs) = backtrack choices |
|
918 |
in prv ([], [], [(!ntrail, length brs, PROVE)], brs) end; |
|
919 |
||
920 |
||
2883 | 921 |
(*Construct an initial branch.*) |
2854 | 922 |
fun initBranch (ts,lim) = |
2894 | 923 |
([(map (fn t => (t,true)) ts, [])], |
924 |
[], add_terms_vars (ts,[]), lim); |
|
2854 | 925 |
|
926 |
||
927 |
(*** Conversion & Skolemization of the Isabelle proof state ***) |
|
928 |
||
929 |
(*Make a list of all the parameters in a subgoal, even if nested*) |
|
930 |
local open Term |
|
931 |
in |
|
932 |
fun discard_foralls (Const("all",_)$Abs(a,T,t)) = discard_foralls t |
|
933 |
| discard_foralls t = t; |
|
934 |
end; |
|
935 |
||
936 |
||
937 |
(*List of variables not appearing as arguments to the given parameter*) |
|
938 |
fun getVars [] i = [] |
|
939 |
| getVars ((_,(v,is))::alist) i = |
|
940 |
if i mem is then getVars alist i |
|
941 |
else v :: getVars alist i; |
|
942 |
||
943 |
||
944 |
(*Conversion of a subgoal: Skolemize all parameters*) |
|
945 |
fun fromSubgoal tsig t = |
|
946 |
let val alist = ref [] |
|
947 |
fun hdvar ((ix,(v,is))::_) = v |
|
948 |
fun from lev t = |
|
949 |
let val (ht,ts) = Term.strip_comb t |
|
950 |
fun apply u = list_comb (u, map (from lev) ts) |
|
951 |
fun bounds [] = [] |
|
952 |
| bounds (Term.Bound i::ts) = |
|
953 |
if i<lev then error"Function Var's argument not a parameter" |
|
954 |
else i-lev :: bounds ts |
|
955 |
| bounds ts = error"Function Var's argument not a bound variable" |
|
956 |
in |
|
957 |
case ht of |
|
958 |
t as Term.Const _ => apply (fromConst tsig t) |
|
959 |
| Term.Free (a,_) => apply (Free a) |
|
960 |
| Term.Bound i => apply (Bound i) |
|
961 |
| Term.Var (ix,_) => |
|
962 |
(case (assoc_string_int(!alist,ix)) of |
|
963 |
None => (alist := (ix, (ref None, bounds ts)) |
|
964 |
:: !alist; |
|
965 |
Var (hdvar(!alist))) |
|
966 |
| Some(v,is) => if is=bounds ts then Var v |
|
967 |
else error("Discrepancy among occurrences of ?" |
|
968 |
^ Syntax.string_of_vname ix)) |
|
969 |
| Term.Abs (a,_,body) => |
|
970 |
if null ts then Abs(a, from (lev+1) body) |
|
971 |
else error "fromSubgoal: argument not in normal form" |
|
972 |
end |
|
973 |
||
974 |
val npars = length (Logic.strip_params t) |
|
975 |
||
976 |
(*Skolemize a subgoal from a proof state*) |
|
977 |
fun skoSubgoal i t = |
|
978 |
if i<npars then |
|
979 |
skoSubgoal (i+1) |
|
980 |
(subst_bound (Skolem (gensym "SG_", getVars (!alist) i), |
|
981 |
t)) |
|
982 |
else t |
|
983 |
||
984 |
in skoSubgoal 0 (from 0 (discard_foralls t)) end; |
|
985 |
||
986 |
||
987 |
(*Tactic using tableau engine and proof reconstruction. |
|
988 |
"lim" is depth limit.*) |
|
989 |
fun depth_tac cs lim i st = |
|
990 |
(fullTrace:=[]; trail := []; ntrail := 0; |
|
991 |
let val {tsig,...} = Sign.rep_sg (#sign (rep_thm st)) |
|
992 |
val skoprem = fromSubgoal tsig (List.nth(prems_of st, i-1)) |
|
993 |
val hyps = strip_imp_prems skoprem |
|
994 |
and concl = strip_imp_concl skoprem |
|
995 |
fun cont (_,choices,tacs) = |
|
996 |
(case Sequence.pull(EVERY' (rev tacs) i st) of |
|
997 |
None => (writeln ("PROOF FAILED for depth " ^ |
|
998 |
Int.toString lim); |
|
999 |
backtrack choices) |
|
1000 |
| cell => Sequence.seqof(fn()=> cell)) |
|
1001 |
in prove (cs, [initBranch (mkGoal concl :: hyps, lim)], cont) |
|
1002 |
end |
|
1003 |
handle Subscript => Sequence.null |
|
1004 |
| PROVE => Sequence.null); |
|
1005 |
||
1006 |
fun blast_tac cs = (DEEPEN (1,20) (depth_tac cs) 0); |
|
1007 |
||
1008 |
fun Blast_tac i = blast_tac (!Data.claset) i; |
|
1009 |
||
2924
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1010 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1011 |
(*** For debugging: these apply the prover to a subgoal and return |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1012 |
the resulting tactics, trace, etc. ***) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1013 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1014 |
(*Translate subgoal i from a proof state*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1015 |
fun trygl cs lim i = |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1016 |
(fullTrace:=[]; trail := []; ntrail := 0; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1017 |
let val st = topthm() |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1018 |
val {tsig,...} = Sign.rep_sg (#sign (rep_thm st)) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1019 |
val skoprem = fromSubgoal tsig (List.nth(prems_of st, i-1)) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1020 |
val hyps = strip_imp_prems skoprem |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1021 |
and concl = strip_imp_concl skoprem |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1022 |
in timeap prove |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1023 |
(cs, [initBranch (mkGoal concl :: hyps, lim)], I) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1024 |
end |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1025 |
handle Subscript => error("There is no subgoal " ^ Int.toString i)); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1026 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1027 |
fun Trygl lim i = trygl (!Data.claset) lim i; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1028 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1029 |
(*Read a string to make an initial, singleton branch*) |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1030 |
fun readGoal sign s = read_cterm sign (s,propT) |> |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1031 |
term_of |> fromTerm (#tsig(Sign.rep_sg sign)) |> |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1032 |
rand |> mkGoal; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1033 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1034 |
fun tryInThy thy lim s = |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1035 |
(fullTrace:=[]; trail := []; ntrail := 0; |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1036 |
timeap prove (!Data.claset, |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1037 |
[initBranch ([readGoal (sign_of thy) s], lim)], |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1038 |
I)); |
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1039 |
|
af506c35b4ed
Control over excessive branching by applying a log2 penalty
paulson
parents:
2894
diff
changeset
|
1040 |
|
2854 | 1041 |
end; |
1042 |