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