author | blanchet |
Thu, 29 Jul 2010 20:02:02 +0200 | |
changeset 38096 | 488b38cd3e06 |
parent 38094 | d01b8119b2e0 |
child 38281 | 601b7972eef2 |
permissions | -rw-r--r-- |
38047 | 1 |
(* Title: HOL/Tools/ATP/atp_systems.ML |
28592 | 2 |
Author: Fabian Immler, TU Muenchen |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
28592 | 4 |
|
36376 | 5 |
Setup for supported ATPs. |
28592 | 6 |
*) |
7 |
||
36376 | 8 |
signature ATP_SYSTEMS = |
28592 | 9 |
sig |
38023 | 10 |
datatype failure = |
11 |
Unprovable | IncompleteUnprovable | CantConnect | TimedOut | |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
12 |
OutOfResources | OldSpass | NoPerl | NoLibwwwPerl | MalformedInput | |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
13 |
MalformedOutput | UnknownError |
38023 | 14 |
|
15 |
type prover_config = |
|
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
16 |
{exec: string * string, |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
17 |
required_execs: (string * string) list, |
38023 | 18 |
arguments: bool -> Time.time -> string, |
19 |
proof_delims: (string * string) list, |
|
20 |
known_failures: (failure * string) list, |
|
21 |
max_new_relevant_facts_per_iter: int, |
|
22 |
prefers_theory_relevant: bool, |
|
23 |
explicit_forall: bool} |
|
24 |
||
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
25 |
val string_for_failure : failure -> string |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
26 |
val known_failure_in_output : |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
27 |
string -> (failure * string) list -> failure option |
38023 | 28 |
val add_prover: string * prover_config -> theory -> theory |
29 |
val get_prover: theory -> string -> prover_config |
|
30 |
val available_atps: theory -> unit |
|
35867 | 31 |
val refresh_systems_on_tptp : unit -> unit |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
32 |
val default_atps_param_value : unit -> string |
35867 | 33 |
val setup : theory -> theory |
28592 | 34 |
end; |
35 |
||
36376 | 36 |
structure ATP_Systems : ATP_SYSTEMS = |
28592 | 37 |
struct |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
38 |
|
38023 | 39 |
(* prover configuration *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
40 |
|
38023 | 41 |
datatype failure = |
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
42 |
Unprovable | IncompleteUnprovable | CantConnect | TimedOut | |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
43 |
OutOfResources | OldSpass | NoPerl | NoLibwwwPerl | MalformedInput | |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
44 |
MalformedOutput | UnknownError |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
45 |
|
32941
72d48e333b77
eliminated extraneous wrapping of public records;
wenzelm
parents:
32936
diff
changeset
|
46 |
type prover_config = |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
47 |
{exec: string * string, |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
48 |
required_execs: (string * string) list, |
37514
b147d01b8ebc
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents:
37509
diff
changeset
|
49 |
arguments: bool -> Time.time -> string, |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
50 |
proof_delims: (string * string) list, |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
51 |
known_failures: (failure * string) list, |
38009
34e1ac9cb71d
shrink the "max_new_relevant_facts_per_iter" fudge factors, now that we count formulas and not clauses
blanchet
parents:
38005
diff
changeset
|
52 |
max_new_relevant_facts_per_iter: int, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
53 |
prefers_theory_relevant: bool, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
54 |
explicit_forall: bool} |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
55 |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
56 |
val missing_message_tail = |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
57 |
" appears to be missing. You will need to install it if you want to run \ |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
58 |
\ATPs remotely." |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
59 |
|
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
60 |
fun string_for_failure Unprovable = "The ATP problem is unprovable." |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
61 |
| string_for_failure IncompleteUnprovable = |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
62 |
"The ATP cannot prove the problem." |
38094 | 63 |
| string_for_failure CantConnect = "Can't connect to remote server." |
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
64 |
| string_for_failure TimedOut = "Timed out." |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
65 |
| string_for_failure OutOfResources = "The ATP ran out of resources." |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
66 |
| string_for_failure OldSpass = |
38096 | 67 |
"Isabelle requires a more recent version of SPASS with support for the \ |
68 |
\TPTP syntax. To install it, download and extract the package \ |
|
69 |
\\"http://isabelle.in.tum.de/dist/contrib/spass-3.7.tar.gz\" and add the \ |
|
70 |
\\"spass-3.7\" directory's absolute path to " ^ |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
71 |
quote (Path.implode (Path.expand (Path.appends |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
72 |
(Path.variable "ISABELLE_HOME_USER" :: |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
73 |
map Path.basic ["etc", "components"])))) ^ |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
74 |
" on a line of its own." |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
75 |
| string_for_failure NoPerl = "Perl" ^ missing_message_tail |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
76 |
| string_for_failure NoLibwwwPerl = |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
77 |
"The Perl module \"libwww-perl\"" ^ missing_message_tail |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
78 |
| string_for_failure MalformedInput = |
38096 | 79 |
"The ATP problem is malformed. Please report this to the Isabelle \ |
80 |
\developers." |
|
81 |
| string_for_failure MalformedOutput = "The ATP output is malformed." |
|
82 |
| string_for_failure UnknownError = "An unknown ATP error occurred." |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
83 |
|
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
84 |
fun known_failure_in_output output = |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
85 |
find_first (fn (_, pattern) => String.isSubstring pattern output) |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
86 |
#> Option.map fst |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
87 |
|
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
88 |
val known_perl_failures = |
38094 | 89 |
[(CantConnect, "HTTP error"), |
90 |
(NoPerl, "env: perl"), |
|
38065 | 91 |
(NoLibwwwPerl, "Can't locate HTTP")] |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
92 |
|
38023 | 93 |
(* named provers *) |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
94 |
|
38023 | 95 |
structure Data = Theory_Data |
96 |
( |
|
97 |
type T = (prover_config * stamp) Symtab.table |
|
98 |
val empty = Symtab.empty |
|
99 |
val extend = I |
|
100 |
fun merge data : T = Symtab.merge (eq_snd op =) data |
|
101 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") |
|
102 |
) |
|
38017
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents:
38015
diff
changeset
|
103 |
|
38023 | 104 |
fun add_prover (name, config) thy = |
105 |
Data.map (Symtab.update_new (name, (config, stamp ()))) thy |
|
106 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") |
|
38017
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents:
38015
diff
changeset
|
107 |
|
38023 | 108 |
fun get_prover thy name = |
109 |
the (Symtab.lookup (Data.get thy) name) |> fst |
|
110 |
handle Option.Option => error ("Unknown ATP: " ^ name ^ ".") |
|
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
111 |
|
38023 | 112 |
fun available_atps thy = |
113 |
priority ("Available ATPs: " ^ |
|
114 |
commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".") |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
115 |
|
38023 | 116 |
fun available_atps thy = |
117 |
priority ("Available ATPs: " ^ |
|
118 |
commas (sort_strings (Symtab.keys (Data.get thy))) ^ ".") |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
119 |
|
36382 | 120 |
fun to_generous_secs time = (Time.toMilliseconds time + 999) div 1000 |
36142
f5e15e9aae10
make Sledgehammer "minimize" output less confusing + round up (not down) time limits to nearest second
blanchet
parents:
36064
diff
changeset
|
121 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
122 |
(* E prover *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
123 |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
124 |
val tstp_proof_delims = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
125 |
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation") |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
126 |
|
35969 | 127 |
val e_config : prover_config = |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
128 |
{exec = ("E_HOME", "eproof"), |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
129 |
required_execs = [], |
37514
b147d01b8ebc
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents:
37509
diff
changeset
|
130 |
arguments = fn _ => fn timeout => |
36382 | 131 |
"--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ |
132 |
string_of_int (to_generous_secs timeout), |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
133 |
proof_delims = [tstp_proof_delims], |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
134 |
known_failures = |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
135 |
[(Unprovable, "SZS status: CounterSatisfiable"), |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
136 |
(Unprovable, "SZS status CounterSatisfiable"), |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
137 |
(TimedOut, "Failure: Resource limit exceeded (time)"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
138 |
(TimedOut, "time limit exceeded"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
139 |
(OutOfResources, |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
140 |
"# Cannot determine problem status within resource limit"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
141 |
(OutOfResources, "SZS status: ResourceOut"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
142 |
(OutOfResources, "SZS status ResourceOut")], |
38090
fe51c098b0ab
fiddle with the fudge factors, to get similar results as before
blanchet
parents:
38065
diff
changeset
|
143 |
max_new_relevant_facts_per_iter = 90 (* FIXME *), |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
144 |
prefers_theory_relevant = false, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
145 |
explicit_forall = false} |
38023 | 146 |
val e = ("e", e_config) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
147 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
148 |
|
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
|
149 |
(* 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
|
150 |
counteracting the presence of "hAPP". *) |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37480
diff
changeset
|
151 |
val spass_config : prover_config = |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
152 |
{exec = ("ISABELLE_ATP", "scripts/spass"), |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
153 |
required_execs = [("SPASS_HOME", "SPASS")], |
37550
fc2f979b9a08
split SPASS time slot between SOS and non-SOS, in case SOS times out
blanchet
parents:
37514
diff
changeset
|
154 |
(* "div 2" accounts for the fact that SPASS is often run twice. *) |
37514
b147d01b8ebc
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents:
37509
diff
changeset
|
155 |
arguments = fn complete => fn timeout => |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
156 |
("-Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ |
37550
fc2f979b9a08
split SPASS time slot between SOS and non-SOS, in case SOS times out
blanchet
parents:
37514
diff
changeset
|
157 |
\-VarWeight=3 -TimeLimit=" ^ |
38094 | 158 |
string_of_int ((to_generous_secs timeout + 1) div 2)) |
37514
b147d01b8ebc
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents:
37509
diff
changeset
|
159 |
|> not complete ? prefix "-SOS=1 ", |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
160 |
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
|
161 |
known_failures = |
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
162 |
known_perl_failures @ |
37413 | 163 |
[(IncompleteUnprovable, "SPASS beiseite: Completion found"), |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
164 |
(TimedOut, "SPASS beiseite: Ran out of time"), |
36965 | 165 |
(OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), |
37413 | 166 |
(MalformedInput, "Undefined symbol"), |
37414
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet
parents:
37413
diff
changeset
|
167 |
(MalformedInput, "Free Variable"), |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
168 |
(OldSpass, "tptp2dfg")], |
38090
fe51c098b0ab
fiddle with the fudge factors, to get similar results as before
blanchet
parents:
38065
diff
changeset
|
169 |
max_new_relevant_facts_per_iter = 35 (* FIXME *), |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
170 |
prefers_theory_relevant = true, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
171 |
explicit_forall = true} |
38023 | 172 |
val spass = ("spass", spass_config) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
173 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
174 |
(* Vampire *) |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
175 |
|
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
176 |
val vampire_config : prover_config = |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
177 |
{exec = ("VAMPIRE_HOME", "vampire"), |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
178 |
required_execs = [], |
37514
b147d01b8ebc
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents:
37509
diff
changeset
|
179 |
arguments = fn _ => fn timeout => |
38033 | 180 |
"--mode casc -t " ^ string_of_int (to_generous_secs timeout) ^ |
181 |
" --input_file ", |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
182 |
proof_delims = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
183 |
[("=========== Refutation ==========", |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
184 |
"======= End of refutation ======="), |
38033 | 185 |
("% SZS output start Refutation", "% SZS output end Refutation"), |
186 |
("% SZS output start Proof", "% SZS output end Proof")], |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
187 |
known_failures = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
188 |
[(Unprovable, "UNPROVABLE"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
189 |
(IncompleteUnprovable, "CANNOT PROVE"), |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
190 |
(TimedOut, "SZS status Timeout"), |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
191 |
(Unprovable, "Satisfiability detected"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
192 |
(OutOfResources, "Refutation not found")], |
38090
fe51c098b0ab
fiddle with the fudge factors, to get similar results as before
blanchet
parents:
38065
diff
changeset
|
193 |
max_new_relevant_facts_per_iter = 50 (* FIXME *), |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
194 |
prefers_theory_relevant = false, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
195 |
explicit_forall = false} |
38023 | 196 |
val vampire = ("vampire", vampire_config) |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
197 |
|
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
198 |
(* Remote prover invocation via SystemOnTPTP *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
199 |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
200 |
val systems = Synchronized.var "atp_systems" ([] : string list) |
31835 | 201 |
|
202 |
fun get_systems () = |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
203 |
case bash_output "\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
204 |
(answer, 0) => split_lines answer |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
205 |
| (answer, _) => |
38065 | 206 |
error (case known_failure_in_output answer known_perl_failures of |
207 |
SOME failure => string_for_failure failure |
|
208 |
| NONE => perhaps (try (unsuffix "\n")) answer ^ ".") |
|
31835 | 209 |
|
35867 | 210 |
fun refresh_systems_on_tptp () = |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
211 |
Synchronized.change systems (fn _ => get_systems ()) |
31835 | 212 |
|
213 |
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
|
214 |
(if null systems then get_systems () else systems) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
215 |
|> `(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
|
216 |
|
32948 | 217 |
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
|
218 |
(case get_system prefix of |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
219 |
NONE => error ("System " ^ quote prefix ^ " not available at SystemOnTPTP.") |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
220 |
| SOME sys => sys); |
31835 | 221 |
|
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
222 |
fun remote_config atp_prefix |
38009
34e1ac9cb71d
shrink the "max_new_relevant_facts_per_iter" fudge factors, now that we count formulas and not clauses
blanchet
parents:
38005
diff
changeset
|
223 |
({proof_delims, known_failures, max_new_relevant_facts_per_iter, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
224 |
prefers_theory_relevant, explicit_forall, ...} : prover_config) |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
225 |
: prover_config = |
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
226 |
{exec = ("ISABELLE_ATP", "scripts/remote_atp"), |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
227 |
required_execs = [], |
37514
b147d01b8ebc
if SPASS fails at finding a proof with the SOS option turned on, turn it off and try again
blanchet
parents:
37509
diff
changeset
|
228 |
arguments = fn _ => fn timeout => |
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
229 |
" -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^ |
36382 | 230 |
the_system atp_prefix, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
231 |
proof_delims = insert (op =) tstp_proof_delims proof_delims, |
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
232 |
known_failures = |
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
233 |
known_failures @ known_perl_failures @ |
38094 | 234 |
[(TimedOut, "says Timeout")], |
38009
34e1ac9cb71d
shrink the "max_new_relevant_facts_per_iter" fudge factors, now that we count formulas and not clauses
blanchet
parents:
38005
diff
changeset
|
235 |
max_new_relevant_facts_per_iter = max_new_relevant_facts_per_iter, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
236 |
prefers_theory_relevant = prefers_theory_relevant, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
237 |
explicit_forall = explicit_forall} |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
238 |
|
38023 | 239 |
val remote_name = prefix "remote_" |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
240 |
|
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
241 |
fun remote_prover (name, config) atp_prefix = |
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
242 |
(remote_name name, remote_config atp_prefix config) |
38023 | 243 |
|
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
244 |
val remote_e = remote_prover e "EP---" |
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
245 |
val remote_vampire = remote_prover vampire "Vampire---9" |
28592 | 246 |
|
38092
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
247 |
fun is_installed ({exec, required_execs, ...} : prover_config) = |
81a003f7de0d
speed up the minimizer by using the time taken for the first iteration as a timeout for the following iterations, and fix a subtle bug in "string_for_failure"
blanchet
parents:
38090
diff
changeset
|
248 |
forall (curry (op <>) "" o getenv o fst) (exec :: required_execs) |
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
249 |
fun maybe_remote (name, config) = |
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
250 |
name |> not (is_installed config) ? remote_name |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
251 |
|
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
252 |
fun default_atps_param_value () = |
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
253 |
space_implode " " ([maybe_remote e] @ |
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
254 |
(if is_installed (snd spass) then [fst spass] else []) @ |
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
255 |
[remote_name (fst vampire)]) |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
256 |
|
38041
3b80d6082131
remove "remote_spass" because there's no way to find out which clauses come from which facts + rename scripts
blanchet
parents:
38033
diff
changeset
|
257 |
val provers = [e, spass, vampire, remote_e, remote_vampire] |
38023 | 258 |
val setup = fold add_prover provers |
35867 | 259 |
|
28592 | 260 |
end; |