author | blanchet |
Fri, 23 Apr 2010 13:16:50 +0200 | |
changeset 36369 | d2cd0d04b8e6 |
parent 36289 | f75b6a3e1450 |
child 36370 | a4f601daa175 |
permissions | -rw-r--r-- |
32327
0971cc0b6a57
src/HOL/Tools/ATP_Manager as separate component, with (almost) everything in one place;
wenzelm
parents:
32257
diff
changeset
|
1 |
(* Title: HOL/Tools/ATP_Manager/atp_wrapper.ML |
28592 | 2 |
Author: Fabian Immler, TU Muenchen |
3 |
||
4 |
Wrapper functions for external ATPs. |
|
5 |
*) |
|
6 |
||
7 |
signature ATP_WRAPPER = |
|
8 |
sig |
|
35867 | 9 |
type prover = ATP_Manager.prover |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
10 |
|
35867 | 11 |
(* hooks for problem files *) |
12 |
val destdir : string Config.T |
|
13 |
val problem_prefix : string Config.T |
|
14 |
val measure_runtime : bool Config.T |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
15 |
|
35867 | 16 |
val refresh_systems_on_tptp : unit -> unit |
17 |
val setup : theory -> theory |
|
28592 | 18 |
end; |
19 |
||
35865 | 20 |
structure ATP_Wrapper : ATP_WRAPPER = |
28592 | 21 |
struct |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
22 |
|
36187 | 23 |
open Sledgehammer_Util |
35969 | 24 |
open Sledgehammer_Fact_Preprocessor |
35865 | 25 |
open Sledgehammer_HOL_Clause |
26 |
open Sledgehammer_Fact_Filter |
|
27 |
open Sledgehammer_Proof_Reconstruct |
|
35867 | 28 |
open ATP_Manager |
35826 | 29 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
30 |
(** generic ATP wrapper **) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
31 |
|
32944
ecc0705174c2
clarified File.platform_path vs. File.shell_path;
wenzelm
parents:
32942
diff
changeset
|
32 |
(* external problem files *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
33 |
|
36001 | 34 |
val (destdir, destdir_setup) = Attrib.config_string "atp_destdir" (K ""); |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
35 |
(*Empty string means create files in Isabelle's temporary files directory.*) |
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
36 |
|
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
37 |
val (problem_prefix, problem_prefix_setup) = |
36001 | 38 |
Attrib.config_string "atp_problem_prefix" (K "prob"); |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
39 |
|
33247 | 40 |
val (measure_runtime, measure_runtime_setup) = |
36001 | 41 |
Attrib.config_bool "atp_measure_runtime" (K false); |
33247 | 42 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
43 |
|
35867 | 44 |
(* prover configuration *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
45 |
|
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
46 |
val remotify = prefix "remote_" |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
47 |
|
32941
72d48e333b77
eliminated extraneous wrapping of public records;
wenzelm
parents:
32936
diff
changeset
|
48 |
type prover_config = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
49 |
{home: string, |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
50 |
executable: string, |
35969 | 51 |
arguments: Time.time -> string, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
52 |
proof_delims: (string * string) list, |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
53 |
known_failures: (string list * string) list, |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
54 |
max_new_clauses: int, |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
55 |
prefers_theory_relevant: bool}; |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
56 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
57 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
58 |
(* basic template *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
59 |
|
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
60 |
fun with_path cleanup after f path = |
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
61 |
Exn.capture f path |
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
62 |
|> tap (fn _ => cleanup path) |
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
63 |
|> Exn.release |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
64 |
|> tap (after path) |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
65 |
|
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
66 |
(* Splits by the first possible of a list of delimiters. *) |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
67 |
fun extract_proof delims output = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
68 |
case pairself (find_first (fn s => String.isSubstring s output)) |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
69 |
(ListPair.unzip delims) of |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
70 |
(SOME begin_delim, SOME end_delim) => |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
71 |
output |> first_field begin_delim |> the |> snd |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
72 |
|> first_field end_delim |> the |> fst |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
73 |
| _ => "" |
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
74 |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
75 |
fun extract_proof_or_failure proof_delims known_failures output = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
76 |
case map_filter |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
77 |
(fn (patterns, message) => |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
78 |
if exists (fn p => String.isSubstring p output) patterns then |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
79 |
SOME message |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
80 |
else |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
81 |
NONE) known_failures of |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
82 |
[] => (case extract_proof proof_delims output of |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
83 |
"" => ("", "Error: The ATP output is malformed.") |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
84 |
| proof => (proof, "")) |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
85 |
| (message :: _) => ("", message) |
35865 | 86 |
|
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
87 |
fun generic_prover overlord get_facts prepare write_file home executable args |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
88 |
proof_delims known_failures name |
36287
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
blanchet
parents:
36286
diff
changeset
|
89 |
({debug, full_types, explicit_apply, isar_proof, modulus, sorts, ...} |
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
blanchet
parents:
36286
diff
changeset
|
90 |
: params) minimize_command |
35969 | 91 |
({subgoal, goal, relevance_override, axiom_clauses, filtered_clauses} |
92 |
: problem) = |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
93 |
let |
31750 | 94 |
(* get clauses and prepare them for writing *) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
95 |
val (ctxt, (chain_ths, th)) = goal; |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
96 |
val thy = ProofContext.theory_of ctxt; |
35865 | 97 |
val chain_ths = map (Thm.put_name_hint chained_hint) chain_ths; |
35969 | 98 |
val goal_cls = #1 (neg_conjecture_clauses ctxt th subgoal); |
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31751
diff
changeset
|
99 |
val the_filtered_clauses = |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
100 |
(case filtered_clauses of |
35969 | 101 |
NONE => get_facts relevance_override goal goal_cls |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
102 |
| SOME fcls => fcls); |
31409
d8537ba165b5
split preparing clauses and writing problemfile;
immler@in.tum.de
parents:
31368
diff
changeset
|
103 |
val the_axiom_clauses = |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
104 |
(case axiom_clauses of |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
105 |
NONE => the_filtered_clauses |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
106 |
| SOME axcls => axcls); |
35969 | 107 |
val (internal_thm_names, clauses) = |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
108 |
prepare goal_cls chain_ths the_axiom_clauses the_filtered_clauses thy; |
31750 | 109 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
110 |
(* path to unique problem file *) |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
111 |
val destdir' = if overlord then getenv "ISABELLE_HOME_USER" |
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
112 |
else Config.get ctxt destdir; |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
113 |
val problem_prefix' = Config.get ctxt problem_prefix; |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
114 |
fun prob_pathname nr = |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
115 |
let |
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
116 |
val probfile = |
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
117 |
Path.basic (problem_prefix' ^ |
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
118 |
(if overlord then "_" ^ name else serial_string ()) |
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
119 |
^ "_" ^ string_of_int nr) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
120 |
in |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
121 |
if destdir' = "" then File.tmp_path probfile |
35570
0e30eef52d85
basic simplification of external_prover signature;
wenzelm
parents:
35010
diff
changeset
|
122 |
else if File.exists (Path.explode destdir') |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36265
diff
changeset
|
123 |
then Path.append (Path.explode destdir') probfile |
36169
27b1cc58715e
store nonmangled names along with mangled type names in Sledgehammer for debugging purposes
blanchet
parents:
36167
diff
changeset
|
124 |
else error ("No such directory: " ^ destdir' ^ ".") |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
125 |
end; |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
126 |
|
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
127 |
val command = Path.explode (home ^ "/" ^ executable) |
31750 | 128 |
(* write out problem file and call prover *) |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
129 |
fun command_line probfile = |
36284
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
blanchet
parents:
36283
diff
changeset
|
130 |
(if Config.get ctxt measure_runtime then |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
131 |
"TIMEFORMAT='%3U'; { time " ^ |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
132 |
space_implode " " [File.shell_path command, args, |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
133 |
File.shell_path probfile] ^ " ; } 2>&1" |
36284
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
blanchet
parents:
36283
diff
changeset
|
134 |
else |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
135 |
space_implode " " ["exec", File.shell_path command, args, |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
136 |
File.shell_path probfile, "2>&1"]) ^ |
36284
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
blanchet
parents:
36283
diff
changeset
|
137 |
(if overlord then |
36286
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
blanchet
parents:
36284
diff
changeset
|
138 |
" | sed 's/,/, /g' \ |
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
blanchet
parents:
36284
diff
changeset
|
139 |
\| sed 's/\\([^!=]\\)\\([=|]\\)\\([^=]\\)/\\1 \\2 \\3/g' \ |
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
blanchet
parents:
36284
diff
changeset
|
140 |
\| sed 's/! =/ !=/g' \ |
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
blanchet
parents:
36284
diff
changeset
|
141 |
\| sed 's/ / /g' | sed 's/| |/||/g' \ |
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
blanchet
parents:
36284
diff
changeset
|
142 |
\| sed 's/ = = =/===/g' \ |
fa6d03d42aab
"remote_e" and "remote_vampire" support TSTP proof output + fix "overlord" mode ATP output postprocessing
blanchet
parents:
36284
diff
changeset
|
143 |
\| sed 's/= = /== /g'" |
36284
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
blanchet
parents:
36283
diff
changeset
|
144 |
else |
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
blanchet
parents:
36283
diff
changeset
|
145 |
"") |
32510 | 146 |
fun split_time s = |
147 |
let |
|
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
148 |
val split = String.tokens (fn c => str c = "\n"); |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
149 |
val (output, t) = s |> split |> split_last |> apfst cat_lines; |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
150 |
fun as_num f = f >> (fst o read_int); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
151 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
152 |
val digit = Scan.one Symbol.is_ascii_digit; |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
153 |
val num3 = as_num (digit ::: digit ::: (digit >> single)); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
154 |
val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
155 |
val as_time = the_default 0 o Scan.read Symbol.stopper time o explode; |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
156 |
in (output, as_time t) end; |
33247 | 157 |
fun split_time' s = |
158 |
if Config.get ctxt measure_runtime then split_time s else (s, 0) |
|
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
159 |
fun run_on probfile = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
160 |
if File.exists command then |
36235
61159615a0c5
added "explicit_apply" option to Sledgehammer, to control whether an explicit apply function should be used as much or as little as possible (replaces a previous global variable)
blanchet
parents:
36231
diff
changeset
|
161 |
write_file full_types explicit_apply probfile clauses |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
162 |
|> pair (apfst split_time' (bash_output (command_line probfile))) |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
163 |
else error ("Bad executable: " ^ Path.implode command ^ "."); |
28592 | 164 |
|
36167
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
36143
diff
changeset
|
165 |
(* If the problem file has not been exported, remove it; otherwise, export |
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
36143
diff
changeset
|
166 |
the proof file too. *) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
167 |
fun cleanup probfile = if destdir' = "" then try File.rm probfile else NONE; |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
168 |
fun export probfile (((output, _), _), _) = |
36187 | 169 |
if destdir' = "" then |
170 |
() |
|
171 |
else |
|
172 |
File.write (Path.explode (Path.implode probfile ^ "_proof")) |
|
36282
9a7c5b86a105
generate command-line in addition to timestamp in ATP output file, for debugging purposes
blanchet
parents:
36281
diff
changeset
|
173 |
((if overlord then |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
174 |
"% " ^ command_line probfile ^ "\n% " ^ timestamp () ^ |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
175 |
"\n" |
36282
9a7c5b86a105
generate command-line in addition to timestamp in ATP output file, for debugging purposes
blanchet
parents:
36281
diff
changeset
|
176 |
else |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
177 |
"") ^ output) |
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32091
diff
changeset
|
178 |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
179 |
val (((output, atp_run_time_in_msecs), rc), _) = |
35570
0e30eef52d85
basic simplification of external_prover signature;
wenzelm
parents:
35010
diff
changeset
|
180 |
with_path cleanup export run_on (prob_pathname subgoal); |
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
181 |
|
36167
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
36143
diff
changeset
|
182 |
(* Check for success and print out some information on failure. *) |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
183 |
val (proof, failure) = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
184 |
extract_proof_or_failure proof_delims known_failures output |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
185 |
val success = (rc = 0 andalso failure = "") |
35969 | 186 |
val (message, relevant_thm_names) = |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
187 |
if success then |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
188 |
proof_text isar_proof debug modulus sorts ctxt |
36287
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
blanchet
parents:
36286
diff
changeset
|
189 |
(minimize_command, proof, internal_thm_names, th, subgoal) |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
190 |
else if failure <> "" then |
36284
0e24322474a4
postprocess ATP output in "overlord" mode to make it more readable
blanchet
parents:
36283
diff
changeset
|
191 |
(failure ^ "\n", []) |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
192 |
else |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
193 |
("Unknown ATP error: " ^ output ^ ".\n", []) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
194 |
in |
36231
bede2d49ba3b
get rid of "conjecture_pos", which is no longer necessary now that it's Metis's job, not Sledgehammer's, to report inconsistent contexts
blanchet
parents:
36223
diff
changeset
|
195 |
{success = success, message = message, |
35969 | 196 |
relevant_thm_names = relevant_thm_names, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
197 |
atp_run_time_in_msecs = atp_run_time_in_msecs, output = output, |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
198 |
proof = proof, internal_thm_names = internal_thm_names, |
35969 | 199 |
filtered_clauses = the_filtered_clauses} |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
200 |
end; |
28592 | 201 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
202 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
203 |
(* generic TPTP-based provers *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
204 |
|
35865 | 205 |
fun generic_tptp_prover |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
206 |
(name, {home, executable, arguments, proof_delims, known_failures, |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
207 |
max_new_clauses, prefers_theory_relevant}) |
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
208 |
(params as {debug, overlord, respect_no_atp, relevance_threshold, |
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
209 |
convergence, theory_relevant, higher_order, follow_defs, |
36287
96f45c5ffb36
if Isar proof reconstruction is not supported, tell the user so they don't wonder why their "isar_proof" option did nothing
blanchet
parents:
36286
diff
changeset
|
210 |
isar_proof, ...}) |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36265
diff
changeset
|
211 |
minimize_command timeout = |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
212 |
generic_prover overlord |
36058
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
blanchet
parents:
36001
diff
changeset
|
213 |
(get_relevant_facts respect_no_atp relevance_threshold convergence |
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
blanchet
parents:
36001
diff
changeset
|
214 |
higher_order follow_defs max_new_clauses |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36219
diff
changeset
|
215 |
(the_default prefers_theory_relevant theory_relevant)) |
36222
0e3e49bd658d
don't use readable names if proof reconstruction is needed, because it uses the structure of names
blanchet
parents:
36220
diff
changeset
|
216 |
(prepare_clauses higher_order false) |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
217 |
(write_tptp_file (debug andalso overlord andalso not isar_proof)) home |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
218 |
executable (arguments timeout) proof_delims known_failures name params |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
219 |
minimize_command |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
220 |
|
35969 | 221 |
fun tptp_prover name p = (name, generic_tptp_prover (name, p)); |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
222 |
|
32941
72d48e333b77
eliminated extraneous wrapping of public records;
wenzelm
parents:
32936
diff
changeset
|
223 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
224 |
(** common provers **) |
28592 | 225 |
|
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36064
diff
changeset
|
226 |
fun generous_to_secs time = (Time.toMilliseconds time + 999) div 1000 |
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36064
diff
changeset
|
227 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
228 |
(* Vampire *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
229 |
|
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
230 |
(* Vampire requires an explicit time limit. *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
231 |
|
35969 | 232 |
val vampire_config : prover_config = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
233 |
{home = getenv "VAMPIRE_HOME", |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
234 |
executable = "vampire", |
35969 | 235 |
arguments = (fn timeout => "--output_syntax tptp --mode casc -t " ^ |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36064
diff
changeset
|
236 |
string_of_int (generous_to_secs timeout)), |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
237 |
proof_delims = [("=========== Refutation ==========", |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
238 |
"======= End of refutation =======")], |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
239 |
known_failures = |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
240 |
[(["Satisfiability detected", "CANNOT PROVE"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
241 |
"The ATP problem is unprovable."), |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
242 |
(["Refutation not found"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
243 |
"The ATP failed to determine the problem's status.")], |
35969 | 244 |
max_new_clauses = 60, |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
245 |
prefers_theory_relevant = false} |
35969 | 246 |
val vampire = tptp_prover "vampire" vampire_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
247 |
|
28592 | 248 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
249 |
(* E prover *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
250 |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
251 |
val tstp_proof_delims = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
252 |
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation") |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
253 |
|
35969 | 254 |
val e_config : prover_config = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
255 |
{home = getenv "E_HOME", |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
256 |
executable = "eproof", |
35969 | 257 |
arguments = (fn timeout => "--tstp-in --tstp-out -l5 -xAutoDev \ |
258 |
\-tAutoDev --silent --cpu-limit=" ^ |
|
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36064
diff
changeset
|
259 |
string_of_int (generous_to_secs timeout)), |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
260 |
proof_delims = [tstp_proof_delims], |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
261 |
known_failures = |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
262 |
[(["SZS status: Satisfiable", "SZS status Satisfiable"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
263 |
"The ATP problem is unprovable."), |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
264 |
(["SZS status: ResourceOut", "SZS status ResourceOut"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
265 |
"The ATP ran out of resources."), |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
266 |
(["# Cannot determine problem status"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
267 |
"The ATP failed to determine the problem's status.")], |
35969 | 268 |
max_new_clauses = 100, |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
269 |
prefers_theory_relevant = false} |
35969 | 270 |
val e = tptp_prover "e" e_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
271 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
272 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
273 |
(* SPASS *) |
28592 | 274 |
|
35865 | 275 |
fun generic_dfg_prover |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
276 |
(name, {home, executable, arguments, proof_delims, known_failures, |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
277 |
max_new_clauses, prefers_theory_relevant}) |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
278 |
(params as {overlord, respect_no_atp, relevance_threshold, convergence, |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36219
diff
changeset
|
279 |
theory_relevant, higher_order, follow_defs, ...}) |
36281
dbbf4d5d584d
pass relevant options from "sledgehammer" to "sledgehammer minimize";
blanchet
parents:
36265
diff
changeset
|
280 |
minimize_command timeout = |
36143
6490319b1703
added "overlord" option (to get easy access to output files for debugging) + systematically use "raw_goal" rather than an inconsistent mixture
blanchet
parents:
36142
diff
changeset
|
281 |
generic_prover overlord |
36058
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
blanchet
parents:
36001
diff
changeset
|
282 |
(get_relevant_facts respect_no_atp relevance_threshold convergence |
8256d5a185bd
added "respect_no_atp" and "convergence" options to Sledgehammer;
blanchet
parents:
36001
diff
changeset
|
283 |
higher_order follow_defs max_new_clauses |
36220
f3655a3ae1ab
rename Sledgehammer "theory_const" option to "theory_relevant", now that I understand better what it does
blanchet
parents:
36219
diff
changeset
|
284 |
(the_default prefers_theory_relevant theory_relevant)) |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
285 |
(prepare_clauses higher_order true) write_dfg_file home executable |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
286 |
(arguments timeout) proof_delims known_failures name params |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
287 |
minimize_command |
32869
159309603edc
recovered support for Spass: re-enabled writing problems in DFG format
boehmes
parents:
32864
diff
changeset
|
288 |
|
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
289 |
fun dfg_prover name p = (name, generic_dfg_prover (name, p)) |
32869
159309603edc
recovered support for Spass: re-enabled writing problems in DFG format
boehmes
parents:
32864
diff
changeset
|
290 |
|
36219
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents:
36190
diff
changeset
|
291 |
(* The "-VarWeight=3" option helps the higher-order problems, probably by |
16670b4f0baa
set SPASS option on the command-line, so that it doesn't vanish when moving to TPTP format
blanchet
parents:
36190
diff
changeset
|
292 |
counteracting the presence of "hAPP". *) |
36059 | 293 |
val spass_config : prover_config = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
294 |
{home = getenv "SPASS_HOME", |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
295 |
executable = "SPASS", |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
296 |
arguments = (fn timeout => "-Auto -SOS=1 -PGiven=0 -PProblem=0 -Splits=0" ^ |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
297 |
" -FullRed=0 -DocProof -VarWeight=3 -TimeLimit=" ^ |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
298 |
string_of_int (generous_to_secs timeout)), |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
299 |
proof_delims = [("Here is a proof", "Formulae used in the proof")], |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
300 |
known_failures = |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
301 |
[(["SPASS beiseite: Completion found."], "The ATP problem is unprovable."), |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
302 |
(["SPASS beiseite: Ran out of time."], "The ATP timed out."), |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
303 |
(["SPASS beiseite: Maximal number of loops exceeded."], |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
304 |
"The ATP hit its loop limit.")], |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
305 |
max_new_clauses = 40, |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
306 |
prefers_theory_relevant = true} |
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
307 |
val spass = dfg_prover "spass" spass_config |
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
308 |
|
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
309 |
(* SPASS 3.7 supports both the DFG and the TPTP syntax, whereas SPASS 3.0 |
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
310 |
supports only the DFG syntax. As soon as all Isabelle repository/snapshot |
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
311 |
users have upgraded to 3.7, we can kill "spass" (and all DFG support in |
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
312 |
Sledgehammer) and rename "spass_tptp" "spass". *) |
35969 | 313 |
|
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
314 |
(* FIXME: Change the error message below to point to the Isabelle download |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
315 |
page once the package is there (around the Isabelle2010 release). *) |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
316 |
|
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
317 |
val spass_tptp_config = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
318 |
{home = #home spass_config, |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
319 |
executable = #executable spass_config, |
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
320 |
arguments = prefix "-TPTP " o #arguments spass_config, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
321 |
proof_delims = #proof_delims spass_config, |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
322 |
known_failures = |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
323 |
#known_failures spass_config @ |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
324 |
[(["unrecognized option `-TPTP'", "Unrecognized option TPTP"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
325 |
"Warning: Sledgehammer requires a more recent version of SPASS with \ |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
326 |
\support for the TPTP syntax. To install it, download and untar the \ |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
327 |
\package \"http://isabelle.in.tum.de/~blanchet/spass-3.7.tgz\" and add \ |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
328 |
\the \"spass-3.7\" directory's full path to \"" ^ |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
329 |
Path.implode (Path.expand (Path.appends |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
330 |
(Path.variable "ISABELLE_HOME_USER" :: |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
331 |
map Path.basic ["etc", "components"]))) ^ |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
332 |
"\" on a line of its own.")], |
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
333 |
max_new_clauses = #max_new_clauses spass_config, |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
334 |
prefers_theory_relevant = #prefers_theory_relevant spass_config} |
36264
3c2490917710
added "spass_tptp" prover, which requires SPASS x.y > 3.0;
blanchet
parents:
36235
diff
changeset
|
335 |
val spass_tptp = tptp_prover "spass_tptp" spass_tptp_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
336 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
337 |
(* remote prover invocation via SystemOnTPTP *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
338 |
|
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
339 |
val systems = Synchronized.var "atp_wrapper_systems" ([]: string list); |
31835 | 340 |
|
341 |
fun get_systems () = |
|
342 |
let |
|
35969 | 343 |
val (answer, rc) = bash_output "\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w" |
31835 | 344 |
in |
35969 | 345 |
if rc <> 0 then |
346 |
error ("Failed to get available systems at SystemOnTPTP:\n" ^ answer) |
|
347 |
else |
|
348 |
split_lines answer |
|
31835 | 349 |
end; |
350 |
||
35867 | 351 |
fun refresh_systems_on_tptp () = |
352 |
Synchronized.change systems (fn _ => get_systems ()); |
|
31835 | 353 |
|
354 |
fun get_system prefix = Synchronized.change_result systems (fn systems => |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
355 |
(if null systems then get_systems () else systems) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
356 |
|> `(find_first (String.isPrefix prefix))); |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
357 |
|
32948 | 358 |
fun the_system prefix = |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
359 |
(case get_system prefix of |
35826 | 360 |
NONE => error ("System " ^ quote prefix ^ " not available at SystemOnTPTP") |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
361 |
| SOME sys => sys); |
31835 | 362 |
|
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
363 |
val remote_known_failures = |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
364 |
[(["Remote-script could not extract proof"], |
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
365 |
"Error: The remote ATP proof is ill-formed.")] |
35865 | 366 |
|
36059 | 367 |
fun remote_prover_config prover_prefix args |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
368 |
({proof_delims, known_failures, max_new_clauses, |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
369 |
prefers_theory_relevant, ...} : prover_config) : prover_config = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
370 |
{home = getenv "ISABELLE_ATP_MANAGER", |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
371 |
executable = "SystemOnTPTP", |
35969 | 372 |
arguments = (fn timeout => |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36064
diff
changeset
|
373 |
args ^ " -t " ^ string_of_int (generous_to_secs timeout) ^ " -s " ^ |
35969 | 374 |
the_system prover_prefix), |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
375 |
proof_delims = insert (op =) tstp_proof_delims proof_delims, |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
376 |
known_failures = remote_known_failures @ known_failures, |
35969 | 377 |
max_new_clauses = max_new_clauses, |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
378 |
prefers_theory_relevant = prefers_theory_relevant} |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
379 |
|
35969 | 380 |
val remote_vampire = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
381 |
tptp_prover (remotify (fst vampire)) |
36059 | 382 |
(remote_prover_config "Vampire---9" "" vampire_config) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
383 |
|
35969 | 384 |
val remote_e = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
385 |
tptp_prover (remotify (fst e)) |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
386 |
(remote_prover_config "EP---" "" e_config) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
387 |
|
35969 | 388 |
val remote_spass = |
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
389 |
tptp_prover (remotify (fst spass)) |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
390 |
(remote_prover_config "SPASS---" "-x" spass_config) |
28592 | 391 |
|
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
392 |
val provers = |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
393 |
[spass, spass_tptp, vampire, e, remote_vampire, remote_spass, remote_e] |
35867 | 394 |
val prover_setup = fold add_prover provers |
395 |
||
396 |
val setup = |
|
397 |
destdir_setup |
|
398 |
#> problem_prefix_setup |
|
399 |
#> measure_runtime_setup |
|
400 |
#> prover_setup; |
|
401 |
||
36289
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
402 |
fun maybe_remote (name, _) ({home, ...} : prover_config) = |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
403 |
name |> home = "" ? remotify |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
404 |
|
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
405 |
val _ = atps := ([maybe_remote e e_config, maybe_remote spass spass_config, |
f75b6a3e1450
set "atps" reference's default value to "(remote_)e (remote_)spass (remote_)vampire", based on what is installed
blanchet
parents:
36287
diff
changeset
|
406 |
remotify (fst vampire)] |> space_implode " ") |
28592 | 407 |
end; |