| author | nipkow | 
| Tue, 24 Mar 2009 14:09:24 +0100 | |
| changeset 30707 | b0391b9b7103 | 
| parent 30542 | eb720644facd | 
| child 30874 | 34927a1e0ae8 | 
| 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: 
30535 
diff
changeset
 | 
13  | 
(thm * (string * int)) list ->  | 
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
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: 
30535 
diff
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: 
28592 
diff
changeset
 | 
17  | 
AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
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: 
28592 
diff
changeset
 | 
19  | 
val tptp_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
20  | 
val tptp_prover: Path.T * string -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
21  | 
val full_prover_opts: int -> bool -> Path.T * string -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
22  | 
val full_prover: Path.T * string -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
23  | 
val vampire_opts: int -> bool -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
24  | 
val vampire: AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
25  | 
val vampire_opts_full: int -> bool -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
26  | 
val vampire_full: AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
27  | 
val eprover_opts: int -> bool -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
28  | 
val eprover: AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
29  | 
val eprover_opts_full: int -> bool -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
30  | 
val eprover_full: AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
31  | 
val spass_opts: int -> bool -> AtpManager.prover  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
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: 
28592 
diff
changeset
 | 
39  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
40  | 
(** generic ATP wrapper **)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
41  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
42  | 
(* global hooks for writing problemfiles *)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
43  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
44  | 
val destdir = ref ""; (*Empty means write files to /tmp*)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
45  | 
val problem_name = ref "prob";  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
46  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
47  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
48  | 
(* basic template *)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
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: 
28592 
diff
changeset
 | 
51  | 
let  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
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: 
28592 
diff
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: 
28592 
diff
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: 
30535 
diff
changeset
 | 
65  | 
val thy = ProofContext.theory_of ctxt  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
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: 
30535 
diff
changeset
 | 
67  | 
val probfile = prob_pathname subgoalno  | 
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
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: 
28592 
diff
changeset
 | 
70  | 
val cmdline =  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
71  | 
if File.exists cmd then File.shell_path cmd ^ " " ^ args  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
72  | 
      else error ("Bad executable: " ^ Path.implode cmd)
 | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
73  | 
val (proof, rc) = system_out (cmdline ^ " " ^ fname)  | 
| 28592 | 74  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
75  | 
(* remove *temporary* files *)  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
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: 
28592 
diff
changeset
 | 
81  | 
val message =  | 
| 30535 | 82  | 
if isSome failure then "Proof attempt failed."  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
83  | 
else if rc <> 0 then "Proof attempt failed: " ^ proof  | 
| 30537 | 84  | 
else "Try this command: " ^ produce_answer (proof, thm_names, ctxt, th, subgoalno)  | 
| 30535 | 85  | 
|
| 
30015
 
1baeda435aa6
detailed information on atp-failure via Output.debug
 
immler@in.tum.de 
parents: 
29597 
diff
changeset
 | 
86  | 
val _ = if isSome failure  | 
| 
 
1baeda435aa6
detailed information on atp-failure via Output.debug
 
immler@in.tum.de 
parents: 
29597 
diff
changeset
 | 
87  | 
then Output.debug (fn () => "Sledgehammer failure: " ^ the failure ^ "\nOutput: " ^ proof)  | 
| 
 
1baeda435aa6
detailed information on atp-failure via Output.debug
 
immler@in.tum.de 
parents: 
29597 
diff
changeset
 | 
88  | 
else ()  | 
| 
 
1baeda435aa6
detailed information on atp-failure via Output.debug
 
immler@in.tum.de 
parents: 
29597 
diff
changeset
 | 
89  | 
val _ = if rc <> 0  | 
| 
 
1baeda435aa6
detailed information on atp-failure via Output.debug
 
immler@in.tum.de 
parents: 
29597 
diff
changeset
 | 
90  | 
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: 
29597 
diff
changeset
 | 
91  | 
else ()  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
92  | 
in (success, message) end;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
93  | 
|
| 28592 | 94  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
95  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
96  | 
(** common provers **)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
97  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
98  | 
(* generic TPTP-based provers *)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
99  | 
|
| 30537 | 100  | 
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: 
28592 
diff
changeset
 | 
101  | 
external_prover  | 
| 30537 | 102  | 
(ResAtp.get_relevant max_new theory_const goal n)  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
103  | 
(ResAtp.write_problem_file false)  | 
| 28592 | 104  | 
command  | 
| 29590 | 105  | 
ResReconstruct.find_failure_e_vamp_spass  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
106  | 
(if full then ResReconstruct.structured_proof else ResReconstruct.lemma_list_tstp)  | 
| 30537 | 107  | 
timeout n goal;  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
108  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
109  | 
(*arbitrary ATP with TPTP input/output and problemfile as last argument*)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
110  | 
fun tptp_prover_opts max_new theory_const =  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
111  | 
tptp_prover_opts_full max_new theory_const false;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
112  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
113  | 
val tptp_prover = tptp_prover_opts 60 true;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
114  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
115  | 
(*for structured proofs: prover must support TSTP*)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
116  | 
fun full_prover_opts max_new theory_const =  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
117  | 
tptp_prover_opts_full max_new theory_const true;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
118  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
119  | 
val full_prover = full_prover_opts 60 true;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
120  | 
|
| 28592 | 121  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
122  | 
(* Vampire *)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
123  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
124  | 
(*NB: Vampire does not work without explicit timelimit*)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
125  | 
|
| 29593 | 126  | 
fun vampire_opts max_new theory_const timeout = tptp_prover_opts  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
127  | 
max_new theory_const  | 
| 29593 | 128  | 
(Path.explode "$VAMPIRE_HOME/vampire",  | 
129  | 
               ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
 | 
|
130  | 
timeout;  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
131  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
132  | 
val vampire = vampire_opts 60 false;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
133  | 
|
| 29593 | 134  | 
fun vampire_opts_full max_new theory_const timeout = full_prover_opts  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
135  | 
max_new theory_const  | 
| 29593 | 136  | 
(Path.explode "$VAMPIRE_HOME/vampire",  | 
137  | 
               ("--output_syntax tptp --mode casc -t " ^ string_of_int timeout))
 | 
|
138  | 
timeout;  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
139  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
140  | 
val vampire_full = vampire_opts 60 false;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
141  | 
|
| 28592 | 142  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
143  | 
(* E prover *)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
144  | 
|
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
145  | 
fun eprover_opts max_new theory_const timeout = tptp_prover_opts  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
146  | 
max_new theory_const  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
147  | 
(Path.explode "$E_HOME/eproof",  | 
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
148  | 
"--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: 
30535 
diff
changeset
 | 
149  | 
timeout;  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
150  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
151  | 
val eprover = eprover_opts 100 false;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
152  | 
|
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
153  | 
fun eprover_opts_full max_new theory_const timeout = full_prover_opts  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
154  | 
max_new theory_const  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
155  | 
(Path.explode "$E_HOME/eproof",  | 
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
156  | 
"--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: 
30535 
diff
changeset
 | 
157  | 
timeout;  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
158  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
159  | 
val eprover_full = eprover_opts_full 100 false;  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
160  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
161  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
162  | 
(* SPASS *)  | 
| 28592 | 163  | 
|
| 30537 | 164  | 
fun spass_opts max_new theory_const timeout n goal = external_prover  | 
165  | 
(ResAtp.get_relevant max_new theory_const goal n)  | 
|
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
166  | 
(ResAtp.write_problem_file true)  | 
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
167  | 
(Path.explode "$SPASS_HOME/SPASS",  | 
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
168  | 
"-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof -TimeLimit=" ^ string_of_int timeout)  | 
| 29590 | 169  | 
ResReconstruct.find_failure_e_vamp_spass  | 
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
170  | 
ResReconstruct.lemma_list_dfg  | 
| 30537 | 171  | 
timeout n goal;  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
172  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
173  | 
val spass = spass_opts 40 true;  | 
| 28592 | 174  | 
|
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
175  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
176  | 
(* remote prover invocation via SystemOnTPTP *)  | 
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
177  | 
|
| 29593 | 178  | 
fun remote_prover_opts max_new theory_const args timeout =  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
179  | 
tptp_prover_opts max_new theory_const  | 
| 
30542
 
eb720644facd
have remote script interrupted like the other provers
 
immler@in.tum.de 
parents: 
30537 
diff
changeset
 | 
180  | 
(Path.explode "$ISABELLE_HOME/contrib/SystemOnTPTP/remote", args ^ " -t " ^ string_of_int timeout)  | 
| 29593 | 181  | 
timeout;  | 
| 
28596
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
182  | 
|
| 
 
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
 
wenzelm 
parents: 
28592 
diff
changeset
 | 
183  | 
val remote_prover = remote_prover_opts 60 false;  | 
| 28592 | 184  | 
|
185  | 
end;  | 
|
| 
30536
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
186  | 
|
| 
 
07b4f050e4df
split relevance-filter and writing of problem-files;
 
immler@in.tum.de 
parents: 
30535 
diff
changeset
 | 
187  |