author | aspinall |
Fri, 30 Sep 2005 18:18:34 +0200 | |
changeset 17740 | fc385ce6187d |
parent 17717 | 7c6a96cbc966 |
child 17746 | af59c748371d |
permissions | -rw-r--r-- |
15608 | 1 |
(* Author: Jia Meng, Cambridge University Computer Laboratory |
2 |
ID: $Id$ |
|
3 |
Copyright 2004 University of Cambridge |
|
15347 | 4 |
|
5 |
ATPs with TPTP format input. |
|
6 |
*) |
|
15452 | 7 |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
8 |
signature RES_ATP = |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
9 |
sig |
17306 | 10 |
val prover: string ref |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
11 |
val custom_spass: string list ref |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
12 |
val destdir: string ref |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
13 |
val hook_count: int ref |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
14 |
val problem_name: string ref |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
15 |
val time_limit: int ref |
15347 | 16 |
end; |
17 |
||
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
18 |
structure ResAtp: RES_ATP = |
15347 | 19 |
struct |
20 |
||
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
21 |
val call_atp = ref false; |
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
22 |
val hook_count = ref 0; |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
23 |
val time_limit = ref 60; |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
24 |
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17317
diff
changeset
|
25 |
val prover = ref "E"; (* use E as the default prover *) |
17305
6cef3aedd661
axioms now included in tptp files, no /bin/cat and various tidying
paulson
parents:
17235
diff
changeset
|
26 |
val custom_spass = (*specialized options for SPASS*) |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
27 |
ref ["-Auto=0","-FullRed=0","-IORe","-IOFc","-RTaut","-RFSub","-RBSub"]; |
15347 | 28 |
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
29 |
val destdir = ref ""; (*Empty means write files to /tmp*) |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
30 |
val problem_name = ref "prob"; |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
31 |
|
17717 | 32 |
fun probfile_nosuffix _ = |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
33 |
if !destdir = "" then File.platform_path (File.tmp_path (Path.basic (!problem_name))) |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
34 |
else if File.exists (File.unpack_platform_path (!destdir)) |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
35 |
then !destdir ^ "/" ^ !problem_name |
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
36 |
else error ("No such directory: " ^ !destdir); |
15644 | 37 |
|
17717 | 38 |
fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n; |
39 |
||
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
40 |
|
15347 | 41 |
(**** for Isabelle/ML interface ****) |
42 |
||
16897 | 43 |
(*Remove unwanted characters such as ? and newline from the textural |
44 |
representation of a theorem (surely they don't need to be produced in |
|
45 |
the first place?) *) |
|
15608 | 46 |
|
16897 | 47 |
fun is_proof_char ch = (#" " <= ch andalso ch <= #"~" andalso ch <> #"?"); |
48 |
||
49 |
val proofstring = |
|
50 |
String.translate (fn c => if is_proof_char c then str c else ""); |
|
15608 | 51 |
|
15452 | 52 |
|
15347 | 53 |
(**** For running in Isar ****) |
54 |
||
15608 | 55 |
(* same function as that in res_axioms.ML *) |
56 |
fun repeat_RS thm1 thm2 = |
|
57 |
let val thm1' = thm1 RS thm2 handle THM _ => thm1 |
|
58 |
in |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
59 |
if eq_thm(thm1,thm1') then thm1' else (repeat_RS thm1' thm2) |
15608 | 60 |
end; |
61 |
||
62 |
(* a special version of repeat_RS *) |
|
17717 | 63 |
fun repeat_someI_ex th = repeat_RS th someI_ex; |
15608 | 64 |
|
16925
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16904
diff
changeset
|
65 |
|
17502 | 66 |
(* write out a subgoal as tptp clauses to the file "xxxx_N"*) |
17717 | 67 |
fun tptp_inputs_tfrees thms pf n (axclauses,classrel_clauses,arity_clauses) = |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
68 |
let |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
69 |
val clss = map (ResClause.make_conjecture_clause_thm) thms |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
70 |
val (tptp_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2tptp clss) |
17422 | 71 |
val tfree_clss = map ResClause.tfree_clause (ResLib.flat_noDup tfree_litss) |
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
72 |
val classrel_cls = map ResClause.tptp_classrelClause classrel_clauses |
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
73 |
val arity_cls = map ResClause.tptp_arity_clause arity_clauses |
17717 | 74 |
val out = TextIO.openOut(pf n) |
15608 | 75 |
in |
17305
6cef3aedd661
axioms now included in tptp files, no /bin/cat and various tidying
paulson
parents:
17235
diff
changeset
|
76 |
ResLib.writeln_strs out (List.concat (map ResClause.tptp_clause axclauses)); |
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
77 |
ResLib.writeln_strs out (tfree_clss @ tptp_clss @ classrel_cls @ arity_cls); |
17717 | 78 |
TextIO.closeOut out |
15608 | 79 |
end; |
15452 | 80 |
|
17502 | 81 |
(* write out a subgoal in DFG format to the file "xxxx_N"*) |
17717 | 82 |
fun dfg_inputs_tfrees thms pf n (axclauses,classrel_clauses,arity_clauses) = |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
83 |
let val clss = map (ResClause.make_conjecture_clause_thm) thms |
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
84 |
(*FIXME: classrel_clauses and arity_clauses*) |
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
85 |
val probN = ResClause.clauses2dfg clss (!problem_name ^ "_" ^ Int.toString n) |
17422 | 86 |
axclauses [] [] [] |
17717 | 87 |
val out = TextIO.openOut(pf n) |
16767 | 88 |
in |
17717 | 89 |
ResLib.writeln_strs out [probN]; TextIO.closeOut out |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
90 |
end; |
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
91 |
|
16357 | 92 |
|
15644 | 93 |
(*********************************************************************) |
17306 | 94 |
(* call prover with settings and problem file for the current subgoal *) |
15644 | 95 |
(*********************************************************************) |
16357 | 96 |
(* now passing in list of skolemized thms and list of sgterms to go with them *) |
17422 | 97 |
fun watcher_call_provers sign sg_terms (childin, childout,pid) = |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
98 |
let |
17422 | 99 |
fun make_atp_list [] n = [] |
17717 | 100 |
| make_atp_list (sg_term::xs) n = |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
101 |
let |
16897 | 102 |
val goalstring = proofstring (Sign.string_of_term sign sg_term) |
16904 | 103 |
val _ = debug ("goalstring in make_atp_lists is " ^ goalstring) |
17717 | 104 |
val probfile = prob_pathname n |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
105 |
val time = Int.toString (!time_limit) |
17502 | 106 |
val _ = debug ("problem file in watcher_call_provers is " ^ probfile) |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
107 |
in |
17317
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
paulson
parents:
17306
diff
changeset
|
108 |
(*Avoid command arguments containing spaces: Poly/ML and SML/NJ |
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
paulson
parents:
17306
diff
changeset
|
109 |
versions of Unix.execute treat them differently!*) |
17306 | 110 |
if !prover = "spass" |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
111 |
then |
17306 | 112 |
let val optionline = |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
113 |
if !AtpCommunication.reconstruct |
17306 | 114 |
(*Proof reconstruction works for only a limited set of |
115 |
inference rules*) |
|
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
116 |
then space_implode "%" (!custom_spass) ^ |
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
117 |
"%-DocProof%-TimeLimit=" ^ time |
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
118 |
else "-DocProof%-SOS%-FullRed=0%-TimeLimit=" ^ time (*Auto mode*) |
16904 | 119 |
val _ = debug ("SPASS option string is " ^ optionline) |
16897 | 120 |
val _ = ResLib.helper_path "SPASS_HOME" "SPASS" |
121 |
(*We've checked that SPASS is there for ATP/spassshell to run.*) |
|
122 |
in |
|
17422 | 123 |
([("spass", goalstring, |
16897 | 124 |
getenv "ISABELLE_HOME" ^ "/src/HOL/Tools/ATP/spassshell", |
17422 | 125 |
optionline, probfile)] @ |
126 |
(make_atp_list xs (n+1))) |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
127 |
end |
17306 | 128 |
else if !prover = "vampire" |
17235
8e55ad29b690
Added ECommunication.ML and modified res_atp.ML, Reconstruction.thy, and
quigley
parents:
17234
diff
changeset
|
129 |
then |
17435 | 130 |
let val vampire = ResLib.helper_path "VAMPIRE_HOME" "vampire" |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
131 |
in |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
132 |
([("vampire", goalstring, vampire, "-m 100000%-t " ^ time, probfile)] @ |
17435 | 133 |
(make_atp_list xs (n+1))) (*BEWARE! spaces in options!*) |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
134 |
end |
17306 | 135 |
else if !prover = "E" |
136 |
then |
|
137 |
let val Eprover = ResLib.helper_path "E_HOME" "eproof" |
|
138 |
in |
|
17422 | 139 |
([("E", goalstring, Eprover, |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
140 |
"--tptp-in%-l5%-xAuto%-tAuto%--cpu-limit=" ^ time, |
17422 | 141 |
probfile)] @ |
142 |
(make_atp_list xs (n+1))) |
|
17306 | 143 |
end |
144 |
else error ("Invalid prover name: " ^ !prover) |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
145 |
end |
15452 | 146 |
|
17422 | 147 |
val atp_list = make_atp_list sg_terms 1 |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
148 |
in |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
149 |
Watcher.callResProvers(childout,atp_list); |
17422 | 150 |
debug "Sent commands to watcher!" |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
151 |
end |
16357 | 152 |
|
17717 | 153 |
(*We write out problem files for each subgoal. Argument pf generates filenames, |
154 |
and allows the suppression of the suffix "_1" in problem-generation mode.*) |
|
155 |
fun write_problem_files pf (ctxt,th) = |
|
156 |
let val prems = Thm.prems_of th |
|
157 |
val (clause_arr, axclauses) = ResClasimp.get_clasimp_lemmas ctxt (hd prems) |
|
158 |
(*FIXME: hack!! need to consider relevance for all prems*) |
|
159 |
val _ = debug ("claset and simprules total clauses = " ^ |
|
160 |
Int.toString (Array.length clause_arr)) |
|
161 |
val thy = ProofContext.theory_of ctxt |
|
162 |
val classrel_clauses = ResTypesSorts.classrel_clauses_thy thy |
|
163 |
val _ = debug ("classrel clauses = " ^ Int.toString (length classrel_clauses)) |
|
164 |
val arity_clauses = ResTypesSorts.arity_clause_thy thy |
|
165 |
val _ = debug ("arity clauses = " ^ Int.toString (length arity_clauses)) |
|
166 |
val write = if !prover = "spass" then dfg_inputs_tfrees else tptp_inputs_tfrees |
|
167 |
fun writenext n = |
|
168 |
if n=0 then () |
|
169 |
else |
|
170 |
(SELECT_GOAL |
|
171 |
(EVERY1 [rtac ccontr, ObjectLogic.atomize_tac, skolemize_tac, |
|
172 |
METAHYPS(fn negs => |
|
173 |
(write (make_clauses negs) pf n |
|
174 |
(axclauses,classrel_clauses,arity_clauses); |
|
175 |
writenext (n-1); |
|
176 |
all_tac))]) n th; |
|
177 |
()) |
|
178 |
in writenext (length prems); clause_arr end; |
|
15644 | 179 |
|
17502 | 180 |
val last_watcher_pid = ref (NONE : Posix.Process.pid option); |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
181 |
|
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
182 |
|
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
183 |
(*writes out the current clasimpset to a tptp file; |
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
184 |
turns off xsymbol at start of function, restoring it at end *) |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
185 |
val isar_atp = setmp print_mode [] |
17717 | 186 |
(fn (ctxt, th) => |
187 |
if Thm.no_prems th then () |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
188 |
else |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
189 |
let |
17502 | 190 |
val _ = (case !last_watcher_pid of NONE => () |
191 |
| SOME pid => (*FIXME: should kill ATP processes too; at least they time out*) |
|
192 |
(debug ("Killing old watcher, pid = " ^ |
|
193 |
Int.toString (ResLib.intOfPid pid)); |
|
194 |
Watcher.killWatcher pid)) |
|
195 |
handle OS.SysErr _ => debug "Attempt to kill watcher failed"; |
|
17717 | 196 |
val clause_arr = write_problem_files prob_pathname (ctxt,th) |
197 |
val (childin, childout, pid) = Watcher.createWatcher (th, clause_arr) |
|
15608 | 198 |
in |
17502 | 199 |
last_watcher_pid := SOME pid; |
17717 | 200 |
debug ("proof state: " ^ string_of_thm th); |
17488
67376a311a2b
further simplification of the Isabelle-ATP linkup
paulson
parents:
17484
diff
changeset
|
201 |
debug ("pid: " ^ Int.toString (ResLib.intOfPid pid)); |
17717 | 202 |
watcher_call_provers (sign_of_thm th) (Thm.prems_of th) (childin, childout, pid) |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
203 |
end); |
15608 | 204 |
|
17422 | 205 |
val isar_atp_writeonly = setmp print_mode [] |
17717 | 206 |
(fn (ctxt,th) => |
207 |
if Thm.no_prems th then () |
|
208 |
else |
|
209 |
let val pf = if Thm.nprems_of th = 1 then probfile_nosuffix |
|
210 |
else prob_pathname |
|
211 |
in ignore (write_problem_files pf (ctxt,th)) end); |
|
15452 | 212 |
|
16357 | 213 |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
214 |
(** the Isar toplevel hook **) |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
215 |
|
17091 | 216 |
val invoke_atp = Toplevel.unknown_proof o Toplevel.keep (fn state => |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
217 |
let |
17091 | 218 |
val proof = Toplevel.proof_of state |
219 |
val (ctxt, (_, goal)) = Proof.get_goal proof |
|
220 |
handle Proof.STATE _ => error "No goal present"; |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
221 |
val thy = ProofContext.theory_of ctxt; |
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
222 |
in |
17091 | 223 |
debug ("initial thm in isar_atp: " ^ |
224 |
Pretty.string_of (ProofContext.pretty_thm ctxt goal)); |
|
225 |
debug ("subgoals in isar_atp: " ^ |
|
226 |
Pretty.string_of (ProofContext.pretty_term ctxt |
|
227 |
(Logic.mk_conjunction_list (Thm.prems_of goal)))); |
|
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
228 |
debug ("number of subgoals in isar_atp: " ^ Int.toString (Thm.nprems_of goal)); |
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
229 |
debug ("current theory: " ^ Context.theory_name thy); |
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
230 |
hook_count := !hook_count +1; |
17717 | 231 |
debug ("in hook for time: " ^ Int.toString (!hook_count)); |
16925
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16904
diff
changeset
|
232 |
ResClause.init thy; |
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
233 |
if !destdir = "" andalso !time_limit > 0 then isar_atp (ctxt, goal) |
17502 | 234 |
else isar_atp_writeonly (ctxt, goal) |
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
235 |
end); |
16357 | 236 |
|
17091 | 237 |
val call_atpP = |
238 |
OuterSyntax.improper_command |
|
239 |
"ProofGeneral.call_atp" |
|
240 |
"call automatic theorem provers" |
|
241 |
OuterKeyword.diag |
|
242 |
(Scan.succeed (Toplevel.no_timing o invoke_atp)); |
|
243 |
||
244 |
val _ = OuterSyntax.add_parsers [call_atpP]; |
|
245 |
||
15347 | 246 |
end; |