| author | wenzelm |
| Sat, 07 Jan 2006 23:27:58 +0100 | |
| changeset 18617 | 8928e8722301 |
| parent 18404 | aa27c10a040e |
| child 18675 | 333a73034023 |
| 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 |
| 17849 | 13 |
val helper_path: string -> string -> string |
|
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 |
| 18003 | 16 |
val writeln_strs: TextIO.outstream -> TextIO.vector list -> unit |
| 15347 | 17 |
end; |
18 |
||
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
19 |
structure ResAtp: RES_ATP = |
| 15347 | 20 |
struct |
21 |
||
|
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
22 |
val call_atp = ref false; |
|
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
23 |
val hook_count = ref 0; |
|
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
24 |
val time_limit = ref 60; |
|
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
25 |
|
|
17404
d16c3a62c396
the experimental tagging system, and the usual tidying
paulson
parents:
17317
diff
changeset
|
26 |
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
|
27 |
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
|
28 |
ref ["-Auto=0","-FullRed=0","-IORe","-IOFc","-RTaut","-RFSub","-RBSub"]; |
| 15347 | 29 |
|
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
30 |
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
|
31 |
val problem_name = ref "prob"; |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
32 |
|
| 17819 | 33 |
(*Return the path to a "helper" like SPASS or tptp2X, first checking that |
34 |
it exists. FIXME: modify to use Path primitives and move to some central place.*) |
|
35 |
fun helper_path evar base = |
|
36 |
case getenv evar of |
|
37 |
"" => error ("Isabelle environment variable " ^ evar ^ " not defined")
|
|
38 |
| home => |
|
39 |
let val path = home ^ "/" ^ base |
|
40 |
in if File.exists (File.unpack_platform_path path) then path |
|
41 |
else error ("Could not find the file " ^ path)
|
|
42 |
end; |
|
43 |
||
| 17717 | 44 |
fun probfile_nosuffix _ = |
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
45 |
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
|
46 |
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
|
47 |
then !destdir ^ "/" ^ !problem_name |
|
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
48 |
else error ("No such directory: " ^ !destdir);
|
| 15644 | 49 |
|
| 17717 | 50 |
fun prob_pathname n = probfile_nosuffix n ^ "_" ^ Int.toString n; |
51 |
||
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
52 |
|
| 15347 | 53 |
(**** For running in Isar ****) |
54 |
||
| 17764 | 55 |
fun writeln_strs _ [] = () |
56 |
| writeln_strs out (s::ss) = (TextIO.output (out, s); TextIO.output (out, "\n"); writeln_strs out ss); |
|
|
16925
0fd7b1438d28
simpler variable names, and no types for monomorphic constants
paulson
parents:
16904
diff
changeset
|
57 |
|
| 17502 | 58 |
(* write out a subgoal as tptp clauses to the file "xxxx_N"*) |
| 17888 | 59 |
fun tptp_inputs_tfrees ths pf n (axclauses,classrel_clauses,arity_clauses) = |
| 17775 | 60 |
let |
| 17888 | 61 |
val clss = ResClause.make_conjecture_clauses (map prop_of ths) |
| 17775 | 62 |
val (tptp_clss,tfree_litss) = ListPair.unzip (map ResClause.clause2tptp clss) |
63 |
val tfree_clss = map ResClause.tfree_clause (foldl (op union_string) [] tfree_litss) |
|
64 |
val classrel_cls = map ResClause.tptp_classrelClause classrel_clauses |
|
65 |
val arity_cls = map ResClause.tptp_arity_clause arity_clauses |
|
66 |
val out = TextIO.openOut(pf n) |
|
67 |
in |
|
68 |
writeln_strs out (List.concat (map ResClause.tptp_clause axclauses)); |
|
69 |
writeln_strs out (tfree_clss @ tptp_clss @ classrel_cls @ arity_cls); |
|
70 |
TextIO.closeOut out |
|
71 |
end; |
|
| 15452 | 72 |
|
| 17502 | 73 |
(* write out a subgoal in DFG format to the file "xxxx_N"*) |
| 17888 | 74 |
fun dfg_inputs_tfrees ths pf n (axclauses,classrel_clauses,arity_clauses) = |
75 |
let val clss = ResClause.make_conjecture_clauses (map prop_of ths) |
|
|
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
76 |
(*FIXME: classrel_clauses and arity_clauses*) |
|
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
77 |
val probN = ResClause.clauses2dfg clss (!problem_name ^ "_" ^ Int.toString n) |
| 17422 | 78 |
axclauses [] [] [] |
| 17717 | 79 |
val out = TextIO.openOut(pf n) |
| 16767 | 80 |
in |
| 17764 | 81 |
writeln_strs out [probN]; TextIO.closeOut out |
|
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
82 |
end; |
|
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
83 |
|
| 16357 | 84 |
|
| 17306 | 85 |
(* call prover with settings and problem file for the current subgoal *) |
| 17764 | 86 |
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
|
87 |
let |
| 17422 | 88 |
fun make_atp_list [] n = [] |
| 17717 | 89 |
| 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
|
90 |
let |
| 17717 | 91 |
val probfile = prob_pathname n |
|
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
92 |
val time = Int.toString (!time_limit) |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
93 |
in |
| 17764 | 94 |
debug ("problem file in watcher_call_provers is " ^ probfile);
|
|
17317
3f12de2e2e6e
Isabelle-ATP link: sortable axiom names; no spaces in switches; general tidying
paulson
parents:
17306
diff
changeset
|
95 |
(*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
|
96 |
versions of Unix.execute treat them differently!*) |
| 17764 | 97 |
(*options are separated by Watcher.setting_sep, currently #"%"*) |
| 17306 | 98 |
if !prover = "spass" |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
99 |
then |
| 17306 | 100 |
let val optionline = |
|
17484
f6a225f97f0a
simplification of the Isabelle-ATP code; hooks for batch generation of problems
paulson
parents:
17435
diff
changeset
|
101 |
if !AtpCommunication.reconstruct |
| 17306 | 102 |
(*Proof reconstruction works for only a limited set of |
103 |
inference rules*) |
|
|
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
104 |
then space_implode "%" (!custom_spass) ^ |
|
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
105 |
"%-DocProof%-TimeLimit=" ^ time |
|
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
106 |
else "-DocProof%-SOS%-FullRed=0%-TimeLimit=" ^ time (*Auto mode*) |
| 16904 | 107 |
val _ = debug ("SPASS option string is " ^ optionline)
|
| 17819 | 108 |
val _ = helper_path "SPASS_HOME" "SPASS" |
| 16897 | 109 |
(*We've checked that SPASS is there for ATP/spassshell to run.*) |
110 |
in |
|
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
111 |
([("spass",
|
| 16897 | 112 |
getenv "ISABELLE_HOME" ^ "/src/HOL/Tools/ATP/spassshell", |
| 17422 | 113 |
optionline, probfile)] @ |
114 |
(make_atp_list xs (n+1))) |
|
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
115 |
end |
| 17306 | 116 |
else if !prover = "vampire" |
|
17235
8e55ad29b690
Added ECommunication.ML and modified res_atp.ML, Reconstruction.thy, and
quigley
parents:
17234
diff
changeset
|
117 |
then |
| 17819 | 118 |
let val vampire = helper_path "VAMPIRE_HOME" "vampire" |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
119 |
in |
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
120 |
([("vampire", vampire, "-m 100000%-t " ^ time, probfile)] @
|
| 17435 | 121 |
(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
|
122 |
end |
| 17306 | 123 |
else if !prover = "E" |
124 |
then |
|
| 17819 | 125 |
let val Eprover = helper_path "E_HOME" "eproof" |
| 17306 | 126 |
in |
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
127 |
([("E", Eprover,
|
|
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
128 |
"--tptp-in%-l5%-xAuto%-tAuto%--cpu-limit=" ^ time, |
| 17422 | 129 |
probfile)] @ |
130 |
(make_atp_list xs (n+1))) |
|
| 17306 | 131 |
end |
132 |
else error ("Invalid prover name: " ^ !prover)
|
|
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
133 |
end |
| 15452 | 134 |
|
| 17422 | 135 |
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
|
136 |
in |
|
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
137 |
Watcher.callResProvers(childout,atp_list); |
| 17422 | 138 |
debug "Sent commands to watcher!" |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
139 |
end |
| 16357 | 140 |
|
| 17717 | 141 |
(*We write out problem files for each subgoal. Argument pf generates filenames, |
142 |
and allows the suppression of the suffix "_1" in problem-generation mode.*) |
|
143 |
fun write_problem_files pf (ctxt,th) = |
|
144 |
let val prems = Thm.prems_of th |
|
145 |
val (clause_arr, axclauses) = ResClasimp.get_clasimp_lemmas ctxt (hd prems) |
|
146 |
(*FIXME: hack!! need to consider relevance for all prems*) |
|
147 |
val _ = debug ("claset and simprules total clauses = " ^
|
|
148 |
Int.toString (Array.length clause_arr)) |
|
149 |
val thy = ProofContext.theory_of ctxt |
|
|
18270
27227433cb42
Only output arities and class relations if !ResClause.keep_types is true.
mengj
parents:
18003
diff
changeset
|
150 |
val classrel_clauses = if !ResClause.keep_types then ResClause.classrel_clauses_thy thy else [] |
| 17717 | 151 |
val _ = debug ("classrel clauses = " ^ Int.toString (length classrel_clauses))
|
|
18270
27227433cb42
Only output arities and class relations if !ResClause.keep_types is true.
mengj
parents:
18003
diff
changeset
|
152 |
val arity_clauses = if !ResClause.keep_types then ResClause.arity_clause_thy thy else [] |
| 17717 | 153 |
val _ = debug ("arity clauses = " ^ Int.toString (length arity_clauses))
|
154 |
val write = if !prover = "spass" then dfg_inputs_tfrees else tptp_inputs_tfrees |
|
155 |
fun writenext n = |
|
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
156 |
if n=0 then [] |
| 17717 | 157 |
else |
158 |
(SELECT_GOAL |
|
159 |
(EVERY1 [rtac ccontr, ObjectLogic.atomize_tac, skolemize_tac, |
|
160 |
METAHYPS(fn negs => |
|
161 |
(write (make_clauses negs) pf n |
|
162 |
(axclauses,classrel_clauses,arity_clauses); |
|
163 |
all_tac))]) n th; |
|
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
164 |
pf n :: writenext (n-1)) |
|
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
165 |
in (writenext (length prems), clause_arr) end; |
| 15644 | 166 |
|
| 17775 | 167 |
val last_watcher_pid = ref (NONE : (TextIO.instream * TextIO.outstream * |
168 |
Posix.Process.pid * string list) option); |
|
169 |
||
170 |
fun kill_last_watcher () = |
|
171 |
(case !last_watcher_pid of |
|
172 |
NONE => () |
|
173 |
| SOME (_, childout, pid, files) => |
|
| 17819 | 174 |
(debug ("Killing old watcher, pid = " ^ string_of_pid pid);
|
| 17775 | 175 |
Watcher.killWatcher pid; |
176 |
ignore (map (try OS.FileSys.remove) files))) |
|
177 |
handle OS.SysErr _ => debug "Attempt to kill watcher failed"; |
|
|
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
178 |
|
|
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
179 |
(*writes out the current clasimpset to a tptp file; |
|
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
180 |
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
|
181 |
val isar_atp = setmp print_mode [] |
| 17717 | 182 |
(fn (ctxt, th) => |
183 |
if Thm.no_prems th then () |
|
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
184 |
else |
|
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
185 |
let |
| 17775 | 186 |
val _ = kill_last_watcher() |
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
187 |
val (files,clause_arr) = write_problem_files prob_pathname (ctxt,th) |
| 17717 | 188 |
val (childin, childout, pid) = Watcher.createWatcher (th, clause_arr) |
| 15608 | 189 |
in |
|
17772
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
190 |
last_watcher_pid := SOME (childin, childout, pid, files); |
|
818cec5f82a4
major simplification: removal of the goalstring argument
paulson
parents:
17764
diff
changeset
|
191 |
debug ("problem files: " ^ space_implode ", " files);
|
| 17819 | 192 |
debug ("pid: " ^ string_of_pid pid);
|
| 17717 | 193 |
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
|
194 |
end); |
| 15608 | 195 |
|
| 17422 | 196 |
val isar_atp_writeonly = setmp print_mode [] |
| 17717 | 197 |
(fn (ctxt,th) => |
198 |
if Thm.no_prems th then () |
|
199 |
else |
|
200 |
let val pf = if Thm.nprems_of th = 1 then probfile_nosuffix |
|
201 |
else prob_pathname |
|
202 |
in ignore (write_problem_files pf (ctxt,th)) end); |
|
| 15452 | 203 |
|
| 16357 | 204 |
|
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
205 |
(** the Isar toplevel hook **) |
|
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
206 |
|
| 17091 | 207 |
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
|
208 |
let |
| 17091 | 209 |
val proof = Toplevel.proof_of state |
210 |
val (ctxt, (_, goal)) = Proof.get_goal proof |
|
211 |
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
|
212 |
val thy = ProofContext.theory_of ctxt; |
|
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
213 |
in |
| 17764 | 214 |
debug ("subgoals in isar_atp:\n" ^
|
| 17091 | 215 |
Pretty.string_of (ProofContext.pretty_term ctxt |
216 |
(Logic.mk_conjunction_list (Thm.prems_of goal)))); |
|
|
17525
ae5bb6001afb
tidying, and support for axclass/classrel clauses
paulson
parents:
17502
diff
changeset
|
217 |
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
|
218 |
debug ("current theory: " ^ Context.theory_name thy);
|
|
17150
ce2a1aeb42aa
DFG output now works for untyped rules (ML "ResClause.untyped();")
quigley
parents:
17091
diff
changeset
|
219 |
hook_count := !hook_count +1; |
| 17717 | 220 |
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
|
221 |
ResClause.init thy; |
|
17690
8ba7c3cd24a8
time limit option; fixed bug concerning first line of ATP output
paulson
parents:
17525
diff
changeset
|
222 |
if !destdir = "" andalso !time_limit > 0 then isar_atp (ctxt, goal) |
| 17502 | 223 |
else isar_atp_writeonly (ctxt, goal) |
|
16802
6eeee59dac4c
use Toplevel.print_state_hook instead of adhoc Proof.atp_hook;
wenzelm
parents:
16767
diff
changeset
|
224 |
end); |
| 16357 | 225 |
|
| 17091 | 226 |
val call_atpP = |
| 17746 | 227 |
OuterSyntax.command |
| 17091 | 228 |
"ProofGeneral.call_atp" |
229 |
"call automatic theorem provers" |
|
230 |
OuterKeyword.diag |
|
231 |
(Scan.succeed (Toplevel.no_timing o invoke_atp)); |
|
232 |
||
233 |
val _ = OuterSyntax.add_parsers [call_atpP]; |
|
234 |
||
| 15347 | 235 |
end; |