| author | Christian Urban <urbanc@in.tum.de> | 
| Sun, 26 Apr 2009 00:42:49 +0200 | |
| changeset 30986 | 047fa04a9fe8 | 
| parent 30899 | d394a17d4fdb | 
| child 30979 | 10eb446df3c7 | 
| permissions | -rw-r--r-- | 
| 28592 | 1 | (* Title: HOL/Tools/atp_wrapper.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Fabian Immler, TU Muenchen | |
| 4 | ||
| 5 | Wrapper functions for external ATPs. | |
| 6 | *) | |
| 7 | ||
| 8 | signature ATP_WRAPPER = | |
| 9 | sig | |
| 10 | val destdir: string ref | |
| 11 | val problem_name: string ref | |
| 12 | val external_prover: | |
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 13 | (thm * (string * int)) list -> | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 14 | (Path.T -> thm -> int -> (thm * (string * int)) list -> theory -> string vector) -> | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 15 | Path.T * string -> (string -> string option) -> | 
| 28592 | 16 | (string * string vector * Proof.context * thm * int -> string) -> | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 17 | AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 18 | val tptp_prover_opts_full: int -> bool -> bool -> Path.T * string -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 19 | val tptp_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 20 | val tptp_prover: Path.T * string -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 21 | val full_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 22 | val full_prover: Path.T * string -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 23 | val vampire_opts: int -> bool -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 24 | val vampire: AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 25 | val vampire_opts_full: int -> bool -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 26 | val vampire_full: AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 27 | val eprover_opts: int -> bool -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 28 | val eprover: AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 29 | val eprover_opts_full: int -> bool -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 30 | val eprover_full: AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 31 | val spass_opts: int -> bool -> AtpManager.prover | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 32 | val spass: AtpManager.prover | 
| 29590 | 33 | val remote_prover_opts: int -> bool -> string -> AtpManager.prover | 
| 34 | val remote_prover: string -> AtpManager.prover | |
| 28592 | 35 | end; | 
| 36 | ||
| 37 | structure AtpWrapper: ATP_WRAPPER = | |
| 38 | struct | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 39 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 40 | (** generic ATP wrapper **) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 41 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 42 | (* global hooks for writing problemfiles *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 43 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 44 | val destdir = ref ""; (*Empty means write files to /tmp*) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 45 | val problem_name = ref "prob"; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 46 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 47 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 48 | (* basic template *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 49 | |
| 30537 | 50 | fun external_prover axiom_clauses write_problem_file (cmd, args) find_failure produce_answer timeout subgoalno goal = | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 51 | let | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 52 | (* path to unique problem file *) | 
| 28592 | 53 | val destdir' = ! destdir | 
| 54 | val problem_name' = ! problem_name | |
| 55 | fun prob_pathname nr = | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 56 | let val probfile = Path.basic (problem_name' ^ serial_string () ^ "_" ^ string_of_int nr) | 
| 28592 | 57 | in if destdir' = "" then File.tmp_path probfile | 
| 58 | else if File.exists (Path.explode (destdir')) | |
| 59 | then Path.append (Path.explode (destdir')) probfile | |
| 60 |         else error ("No such directory: " ^ destdir')
 | |
| 61 | end | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 62 | |
| 28592 | 63 | (* write out problem file and call prover *) | 
| 30537 | 64 | val (ctxt, (chain_ths, th)) = goal | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 65 | val thy = ProofContext.theory_of ctxt | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 66 | val chain_ths = map (Thm.put_name_hint ResReconstruct.chained_hint) chain_ths | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 67 | val probfile = prob_pathname subgoalno | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 68 | val fname = File.platform_path probfile | 
| 30537 | 69 | val thm_names = write_problem_file probfile th subgoalno axiom_clauses thy | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 70 | val cmdline = | 
| 30899 
d394a17d4fdb
external_prover: "exec" the command line, in order to preserve the exact process context of the "system" invocation (this recovers interruptibility of E-1.0, which assumes to be the process group leader);
 wenzelm parents: 
30896diff
changeset | 71 | if File.exists cmd then "exec " ^ File.shell_path cmd ^ " " ^ args | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 72 |       else error ("Bad executable: " ^ Path.implode cmd)
 | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 73 | val (proof, rc) = system_out (cmdline ^ " " ^ fname) | 
| 28592 | 74 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 75 | (* remove *temporary* files *) | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 76 | val _ = if destdir' = "" then OS.FileSys.remove fname else () | 
| 29590 | 77 | |
| 78 | (* check for success and print out some information on failure *) | |
| 79 | val failure = find_failure proof | |
| 29597 | 80 | val success = rc = 0 andalso is_none failure | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 81 | val message = | 
| 30896 
ec3f33437fe3
more generic error message, which also covers more fundamental failure;
 wenzelm parents: 
30874diff
changeset | 82 | if is_some failure then "External prover failed." | 
| 
ec3f33437fe3
more generic error message, which also covers more fundamental failure;
 wenzelm parents: 
30874diff
changeset | 83 | else if rc <> 0 then "External prover failed: " ^ proof | 
| 30537 | 84 | else "Try this command: " ^ produce_answer (proof, thm_names, ctxt, th, subgoalno) | 
| 30535 | 85 | |
| 30896 
ec3f33437fe3
more generic error message, which also covers more fundamental failure;
 wenzelm parents: 
30874diff
changeset | 86 | val _ = | 
| 
ec3f33437fe3
more generic error message, which also covers more fundamental failure;
 wenzelm parents: 
30874diff
changeset | 87 | if is_some failure | 
| 30015 
1baeda435aa6
detailed information on atp-failure via Output.debug
 immler@in.tum.de parents: 
29597diff
changeset | 88 | then Output.debug (fn () => "Sledgehammer failure: " ^ the failure ^ "\nOutput: " ^ proof) | 
| 
1baeda435aa6
detailed information on atp-failure via Output.debug
 immler@in.tum.de parents: 
29597diff
changeset | 89 | else () | 
| 30896 
ec3f33437fe3
more generic error message, which also covers more fundamental failure;
 wenzelm parents: 
30874diff
changeset | 90 | val _ = | 
| 
ec3f33437fe3
more generic error message, which also covers more fundamental failure;
 wenzelm parents: 
30874diff
changeset | 91 | if rc <> 0 | 
| 30015 
1baeda435aa6
detailed information on atp-failure via Output.debug
 immler@in.tum.de parents: 
29597diff
changeset | 92 | then Output.debug (fn () => "Sledgehammer exited with return code " ^ string_of_int rc ^ ":\n" ^ proof) | 
| 
1baeda435aa6
detailed information on atp-failure via Output.debug
 immler@in.tum.de parents: 
29597diff
changeset | 93 | else () | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 94 | in (success, message) end; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 95 | |
| 28592 | 96 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 97 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 98 | (** common provers **) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 99 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 100 | (* generic TPTP-based provers *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 101 | |
| 30537 | 102 | fun tptp_prover_opts_full max_new theory_const full command timeout n goal = | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 103 | external_prover | 
| 30537 | 104 | (ResAtp.get_relevant max_new theory_const goal n) | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 105 | (ResAtp.write_problem_file false) | 
| 28592 | 106 | command | 
| 30874 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30542diff
changeset | 107 | ResReconstruct.find_failure | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 108 | (if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list_tstp) | 
| 30537 | 109 | timeout n goal; | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 110 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 111 | (*arbitrary ATP with TPTP input/output and problemfile as last argument*) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 112 | fun tptp_prover_opts max_new theory_const = | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 113 | tptp_prover_opts_full max_new theory_const false; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 114 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 115 | val tptp_prover = tptp_prover_opts 60 true; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 116 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 117 | (*for structured proofs: prover must support TSTP*) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 118 | fun full_prover_opts max_new theory_const = | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 119 | tptp_prover_opts_full max_new theory_const true; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 120 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 121 | val full_prover = full_prover_opts 60 true; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 122 | |
| 28592 | 123 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 124 | (* Vampire *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 125 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 126 | (*NB: Vampire does not work without explicit timelimit*) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 127 | |
| 29593 | 128 | fun vampire_opts max_new theory_const timeout = tptp_prover_opts | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 129 | max_new theory_const | 
| 29593 | 130 | (Path.explode "$VAMPIRE_HOME/vampire", | 
| 131 |                ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
 | |
| 132 | timeout; | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 133 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 134 | val vampire = vampire_opts 60 false; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 135 | |
| 29593 | 136 | fun vampire_opts_full max_new theory_const timeout = full_prover_opts | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 137 | max_new theory_const | 
| 29593 | 138 | (Path.explode "$VAMPIRE_HOME/vampire", | 
| 139 |                ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
 | |
| 140 | timeout; | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 141 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 142 | val vampire_full = vampire_opts 60 false; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 143 | |
| 28592 | 144 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 145 | (* E prover *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 146 | |
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 147 | fun eprover_opts max_new theory_const timeout = tptp_prover_opts | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 148 | max_new theory_const | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 149 | (Path.explode "$E_HOME/eproof", | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 150 | "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout) | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 151 | timeout; | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 152 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 153 | val eprover = eprover_opts 100 false; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 154 | |
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 155 | fun eprover_opts_full max_new theory_const timeout = full_prover_opts | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 156 | max_new theory_const | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 157 | (Path.explode "$E_HOME/eproof", | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 158 | "--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ string_of_int timeout) | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 159 | timeout; | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 160 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 161 | val eprover_full = eprover_opts_full 100 false; | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 162 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 163 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 164 | (* SPASS *) | 
| 28592 | 165 | |
| 30537 | 166 | fun spass_opts max_new theory_const timeout n goal = external_prover | 
| 167 | (ResAtp.get_relevant max_new theory_const goal n) | |
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 168 | (ResAtp.write_problem_file true) | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 169 | (Path.explode "$SPASS_HOME/SPASS", | 
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 170 | "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^ string_of_int timeout) | 
| 30874 
34927a1e0ae8
reverted to explicitly check the presence of a refutation
 immler@in.tum.de parents: 
30542diff
changeset | 171 | ResReconstruct.find_failure | 
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 172 | ResReconstruct.lemma_list_dfg | 
| 30537 | 173 | timeout n goal; | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 174 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 175 | val spass = spass_opts 40 true; | 
| 28592 | 176 | |
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 177 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 178 | (* remote prover invocation via SystemOnTPTP *) | 
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 179 | |
| 29593 | 180 | fun remote_prover_opts max_new theory_const args timeout = | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 181 | tptp_prover_opts max_new theory_const | 
| 30542 
eb720644facd
have remote script interrupted like the other provers
 immler@in.tum.de parents: 
30537diff
changeset | 182 | (Path.explode "$ISABELLE_HOME/contrib/SystemOnTPTP/remote", args ^ " -t " ^ string_of_int timeout) | 
| 29593 | 183 | timeout; | 
| 28596 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 184 | |
| 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 wenzelm parents: 
28592diff
changeset | 185 | val remote_prover = remote_prover_opts 60 false; | 
| 28592 | 186 | |
| 187 | end; | |
| 30536 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 188 | |
| 
07b4f050e4df
split relevance-filter and writing of problem-files;
 immler@in.tum.de parents: 
30535diff
changeset | 189 |