author | blanchet |
Wed, 23 Jun 2010 11:36:03 +0200 | |
changeset 37514 | b147d01b8ebc |
parent 37509 | f39464d971c4 |
child 37550 | fc2f979b9a08 |
permissions | -rw-r--r-- |
36376 | 1 |
(* Title: HOL/Tools/ATP_Manager/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 |
35867 | 10 |
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
|
11 |
|
35867 | 12 |
(* hooks for problem files *) |
36376 | 13 |
val dest_dir : string Config.T |
35867 | 14 |
val problem_prefix : string Config.T |
15 |
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
|
16 |
|
35867 | 17 |
val refresh_systems_on_tptp : unit -> unit |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
18 |
val default_atps_param_value : unit -> string |
35867 | 19 |
val setup : theory -> theory |
28592 | 20 |
end; |
21 |
||
36376 | 22 |
structure ATP_Systems : ATP_SYSTEMS = |
28592 | 23 |
struct |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
24 |
|
36187 | 25 |
open Sledgehammer_Util |
35969 | 26 |
open Sledgehammer_Fact_Preprocessor |
35865 | 27 |
open Sledgehammer_HOL_Clause |
28 |
open Sledgehammer_Fact_Filter |
|
29 |
open Sledgehammer_Proof_Reconstruct |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
30 |
open Sledgehammer_TPTP_Format |
35867 | 31 |
open ATP_Manager |
35826 | 32 |
|
36376 | 33 |
(** generic ATP **) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
34 |
|
32944
ecc0705174c2
clarified File.platform_path vs. File.shell_path;
wenzelm
parents:
32942
diff
changeset
|
35 |
(* external problem files *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
36 |
|
36376 | 37 |
val (dest_dir, dest_dir_setup) = Attrib.config_string "atp_dest_dir" (K ""); |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
38 |
(*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
|
39 |
|
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
40 |
val (problem_prefix, problem_prefix_setup) = |
36001 | 41 |
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
|
42 |
|
33247 | 43 |
val (measure_runtime, measure_runtime_setup) = |
36001 | 44 |
Attrib.config_bool "atp_measure_runtime" (K false); |
33247 | 45 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
46 |
|
35867 | 47 |
(* prover configuration *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
48 |
|
32941
72d48e333b77
eliminated extraneous wrapping of public records;
wenzelm
parents:
32936
diff
changeset
|
49 |
type prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
50 |
{home_var: string, |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
51 |
executable: string, |
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
|
52 |
arguments: bool -> Time.time -> string, |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
53 |
proof_delims: (string * string) list, |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
54 |
known_failures: (failure * string) list, |
36382 | 55 |
max_axiom_clauses: int, |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
56 |
prefers_theory_relevant: bool}; |
28596
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 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
59 |
(* basic template *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
60 |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
61 |
val remotify = prefix "remote_" |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
62 |
|
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
63 |
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
|
64 |
Exn.capture f path |
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
65 |
|> tap (fn _ => cleanup path) |
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
66 |
|> Exn.release |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
67 |
|> tap (after path) |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
68 |
|
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
69 |
(* 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
|
70 |
fun extract_proof delims output = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
71 |
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
|
72 |
(ListPair.unzip delims) of |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
73 |
(SOME begin_delim, SOME end_delim) => |
36549 | 74 |
(output |> first_field begin_delim |> the |> snd |
75 |
|> first_field end_delim |> the |> fst |
|
76 |
|> first_field "\n" |> the |> snd |
|
77 |
handle Option.Option => "") |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
78 |
| _ => "" |
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
79 |
|
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
|
80 |
fun extract_proof_and_outcome complete res_code proof_delims known_failures |
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
|
81 |
output = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
82 |
case map_filter (fn (failure, pattern) => |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
83 |
if String.isSubstring pattern output then SOME failure |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
84 |
else NONE) known_failures of |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
85 |
[] => (case extract_proof proof_delims output of |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
86 |
"" => ("", SOME UnknownError) |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
87 |
| proof => if res_code = 0 then (proof, NONE) |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
88 |
else ("", SOME UnknownError)) |
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
|
89 |
| (failure :: _) => |
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
|
90 |
("", SOME (if failure = IncompleteUnprovable andalso complete then |
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
|
91 |
Unprovable |
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
|
92 |
else |
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
|
93 |
failure)) |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
94 |
|
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
95 |
fun string_for_failure Unprovable = "The ATP problem is unprovable." |
37413 | 96 |
| string_for_failure IncompleteUnprovable = |
97 |
"The ATP cannot prove the problem." |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
98 |
| string_for_failure TimedOut = "Timed out." |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
99 |
| string_for_failure OutOfResources = "The ATP ran out of resources." |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
100 |
| string_for_failure OldSpass = |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36382
diff
changeset
|
101 |
(* FIXME: Change the error message below to point to the Isabelle download |
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36382
diff
changeset
|
102 |
page once the package is there (around the Isabelle2010 release). *) |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
103 |
"Warning: Sledgehammer requires a more recent version of SPASS with \ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
104 |
\support for the TPTP syntax. To install it, download and untar the \ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
105 |
\package \"http://isabelle.in.tum.de/~blanchet/spass-3.7.tgz\" and add the \ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
106 |
\\"spass-3.7\" directory's full path to \"" ^ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
107 |
Path.implode (Path.expand (Path.appends |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
108 |
(Path.variable "ISABELLE_HOME_USER" :: |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
109 |
map Path.basic ["etc", "components"]))) ^ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
110 |
"\" on a line of its own." |
36965 | 111 |
| string_for_failure MalformedInput = |
112 |
"Internal Sledgehammer error: The ATP problem is malformed. Please report \ |
|
113 |
\this to the Isabelle developers." |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
114 |
| string_for_failure MalformedOutput = "Error: The ATP output is malformed." |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
115 |
| string_for_failure UnknownError = "Error: An unknown ATP error occurred." |
35865 | 116 |
|
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
117 |
fun shape_of_clauses _ [] = [] |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
118 |
| shape_of_clauses j ([] :: clauses) = [] :: shape_of_clauses j clauses |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37480
diff
changeset
|
119 |
| shape_of_clauses j ((_ :: lits) :: clauses) = |
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
120 |
let val shape = shape_of_clauses (j + 1) (lits :: clauses) in |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
121 |
(j :: hd shape) :: tl shape |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
122 |
end |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
123 |
|
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
124 |
(* generic TPTP-based provers *) |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
125 |
|
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
126 |
fun generic_tptp_prover |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
127 |
(name, {home_var, executable, arguments, proof_delims, known_failures, |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
128 |
max_axiom_clauses, prefers_theory_relevant}) |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
129 |
({debug, overlord, full_types, explicit_apply, respect_no_atp, |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
130 |
relevance_threshold, relevance_convergence, theory_relevant, |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
131 |
defs_relevant, isar_proof, isar_shrink_factor, ...} : params) |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
132 |
minimize_command timeout |
35969 | 133 |
({subgoal, goal, relevance_override, axiom_clauses, filtered_clauses} |
134 |
: problem) = |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
135 |
let |
31750 | 136 |
(* get clauses and prepare them for writing *) |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37480
diff
changeset
|
137 |
val (ctxt, (_, th)) = goal; |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
138 |
val thy = ProofContext.theory_of ctxt; |
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
139 |
val goal_clss = #1 (neg_conjecture_clauses ctxt th subgoal) |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
140 |
val goal_cls = List.concat goal_clss |
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31751
diff
changeset
|
141 |
val the_filtered_clauses = |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
142 |
(case filtered_clauses of |
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
143 |
NONE => relevant_facts full_types respect_no_atp relevance_threshold |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
144 |
relevance_convergence defs_relevant max_axiom_clauses |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
145 |
(the_default prefers_theory_relevant theory_relevant) |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
146 |
relevance_override goal goal_cls |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
147 |
| SOME fcls => fcls); |
37506
32a1ee39c49b
missing "Unsynchronized" + make exception take a unit
blanchet
parents:
37499
diff
changeset
|
148 |
val the_axiom_clauses = axiom_clauses |> the_default the_filtered_clauses |
35969 | 149 |
val (internal_thm_names, clauses) = |
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
150 |
prepare_clauses full_types goal_cls the_axiom_clauses the_filtered_clauses |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
151 |
thy |
31750 | 152 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
153 |
(* path to unique problem file *) |
36376 | 154 |
val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER" |
155 |
else Config.get ctxt dest_dir; |
|
156 |
val the_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
|
157 |
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
|
158 |
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
|
159 |
val probfile = |
36568
d495d2e1f0a6
in "overlord" mode: ignore problem prefix specified in the .thy file
blanchet
parents:
36552
diff
changeset
|
160 |
Path.basic ((if overlord then "prob_" ^ name |
d495d2e1f0a6
in "overlord" mode: ignore problem prefix specified in the .thy file
blanchet
parents:
36552
diff
changeset
|
161 |
else the_problem_prefix ^ serial_string ()) |
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
|
162 |
^ "_" ^ string_of_int nr) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
163 |
in |
36376 | 164 |
if the_dest_dir = "" then File.tmp_path probfile |
165 |
else if File.exists (Path.explode the_dest_dir) |
|
166 |
then Path.append (Path.explode the_dest_dir) probfile |
|
167 |
else error ("No such directory: " ^ the_dest_dir ^ ".") |
|
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
168 |
end; |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
169 |
|
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
170 |
val home = getenv home_var |
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
|
171 |
val command = Path.explode (home ^ "/" ^ executable) |
31750 | 172 |
(* write out problem file and call prover *) |
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
|
173 |
fun command_line complete probfile = |
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
|
174 |
let |
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
|
175 |
val core = File.shell_path command ^ " " ^ arguments complete timeout ^ |
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
|
176 |
" " ^ File.shell_path probfile |
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
|
177 |
in |
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
|
178 |
(if Config.get ctxt measure_runtime then |
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 |
"TIMEFORMAT='%3U'; { time " ^ core ^ " ; }" |
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
|
180 |
else |
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
|
181 |
"exec " ^ core) ^ " 2>&1" ^ |
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
|
182 |
(if overlord then |
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
|
183 |
" | sed 's/,/, /g' \ |
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
|
184 |
\| sed 's/\\([^!=<]\\)\\([=|]\\)\\([^=>]\\)/\\1 \\2 \\3/g' \ |
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
|
185 |
\| sed 's/ / /g' | sed 's/| |/||/g' \ |
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
|
186 |
\| sed 's/ = = =/===/g' \ |
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
|
187 |
\| sed 's/= = /== /g'" |
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
|
188 |
else |
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
|
189 |
"") |
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
|
190 |
end |
32510 | 191 |
fun split_time s = |
192 |
let |
|
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
193 |
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
|
194 |
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
|
195 |
fun as_num f = f >> (fst o read_int); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
196 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
197 |
val digit = Scan.one Symbol.is_ascii_digit; |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
198 |
val num3 = as_num (digit ::: digit ::: (digit >> single)); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
199 |
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
|
200 |
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
|
201 |
in (output, as_time t) end; |
32458
de6834b20e9e
sledgehammer's temporary files are removed properly (even in case of an exception occurs)
boehmes
parents:
32451
diff
changeset
|
202 |
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
|
203 |
if File.exists command then |
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
|
204 |
let |
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
|
205 |
fun do_run complete = |
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
|
206 |
let |
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
|
207 |
val command = command_line complete probfile |
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
|
208 |
val ((output, msecs), res_code) = |
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
|
209 |
bash_output command |
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
|
210 |
|>> (if overlord then |
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
|
211 |
prefix ("% " ^ command ^ "\n% " ^ timestamp () ^ "\n") |
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
|
212 |
else |
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
|
213 |
I) |
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
|
214 |
|>> (if Config.get ctxt measure_runtime then split_time |
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
|
215 |
else rpair 0) |
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
|
216 |
val (proof, outcome) = |
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
|
217 |
extract_proof_and_outcome complete res_code proof_delims |
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
|
218 |
known_failures output |
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
|
219 |
in (output, msecs, proof, outcome) end |
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
|
220 |
val (pool, conjecture_offset) = |
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
|
221 |
write_tptp_file (debug andalso overlord) full_types explicit_apply |
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
|
222 |
probfile clauses |
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
|
223 |
val conjecture_shape = shape_of_clauses (conjecture_offset + 1) goal_clss |
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
|
224 |
val result = |
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
|
225 |
do_run false |
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
|
226 |
|> (fn (_, msecs0, _, SOME IncompleteUnprovable) => |
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
|
227 |
do_run true |
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 |
|> (fn (output, msecs, proof, outcome) => |
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
|
229 |
(output, msecs0 + msecs, proof, outcome)) |
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
|
230 |
| result => result) |
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
|
231 |
in ((pool, conjecture_shape), result) end |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
232 |
else if home = "" then |
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
233 |
error ("The environment variable " ^ quote home_var ^ " is not set.") |
36382 | 234 |
else |
235 |
error ("Bad executable: " ^ Path.implode command ^ "."); |
|
28592 | 236 |
|
36167
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
36143
diff
changeset
|
237 |
(* 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
|
238 |
the proof file too. *) |
36376 | 239 |
fun cleanup probfile = |
240 |
if the_dest_dir = "" then try File.rm probfile else NONE |
|
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
|
241 |
fun export probfile (_, (output, _, _, _)) = |
36376 | 242 |
if the_dest_dir = "" then |
36187 | 243 |
() |
244 |
else |
|
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
|
245 |
File.write (Path.explode (Path.implode probfile ^ "_proof")) output |
32257
bad5a99c16d8
neg_conjecture_clauses, neg_clausify_tac: proper context, eliminated METAHYPS;
wenzelm
parents:
32091
diff
changeset
|
246 |
|
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
|
247 |
val ((pool, conjecture_shape), (output, msecs, proof, outcome)) = |
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
|
248 |
with_path cleanup export run_on (prob_pathname subgoal) |
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
|
249 |
|
35969 | 250 |
val (message, relevant_thm_names) = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
251 |
case outcome of |
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
252 |
NONE => |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
253 |
proof_text isar_proof |
37480 | 254 |
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape) |
255 |
(full_types, minimize_command, proof, internal_thm_names, th, |
|
256 |
subgoal) |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
257 |
| SOME failure => (string_for_failure failure ^ "\n", []) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
258 |
in |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36382
diff
changeset
|
259 |
{outcome = outcome, message = message, pool = pool, |
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
|
260 |
relevant_thm_names = relevant_thm_names, atp_run_time_in_msecs = msecs, |
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
|
261 |
output = output, proof = proof, internal_thm_names = internal_thm_names, |
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
262 |
conjecture_shape = conjecture_shape, |
35969 | 263 |
filtered_clauses = the_filtered_clauses} |
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
264 |
end |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
265 |
|
35969 | 266 |
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
|
267 |
|
36382 | 268 |
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
|
269 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
270 |
(* E prover *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
271 |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
272 |
val tstp_proof_delims = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
273 |
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation") |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
274 |
|
35969 | 275 |
val e_config : prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
276 |
{home_var = "E_HOME", |
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
|
277 |
executable = "eproof", |
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
|
278 |
arguments = fn _ => fn timeout => |
36382 | 279 |
"--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ |
280 |
string_of_int (to_generous_secs timeout), |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
281 |
proof_delims = [tstp_proof_delims], |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
282 |
known_failures = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
283 |
[(Unprovable, "SZS status: Satisfiable"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
284 |
(Unprovable, "SZS status Satisfiable"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
285 |
(TimedOut, "Failure: Resource limit exceeded (time)"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
286 |
(TimedOut, "time limit exceeded"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
287 |
(OutOfResources, |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
288 |
"# Cannot determine problem status within resource limit"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
289 |
(OutOfResources, "SZS status: ResourceOut"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
290 |
(OutOfResources, "SZS status ResourceOut")], |
36382 | 291 |
max_axiom_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
|
292 |
prefers_theory_relevant = false} |
35969 | 293 |
val e = tptp_prover "e" e_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
294 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
295 |
|
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
|
296 |
(* 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
|
297 |
counteracting the presence of "hAPP". *) |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37480
diff
changeset
|
298 |
val spass_config : prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
299 |
{home_var = "SPASS_HOME", |
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 |
executable = "SPASS", |
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
|
301 |
arguments = fn complete => fn timeout => |
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
|
302 |
("-TPTP -Auto -PGiven=0 -PProblem=0 -Splits=0 -FullRed=0 -DocProof \ |
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
|
303 |
\-VarWeight=3 -TimeLimit=" ^ string_of_int (to_generous_secs timeout)) |
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
|
304 |
|> not complete ? prefix "-SOS=1 ", |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
305 |
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
|
306 |
known_failures = |
37413 | 307 |
[(IncompleteUnprovable, "SPASS beiseite: Completion found"), |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
308 |
(TimedOut, "SPASS beiseite: Ran out of time"), |
36965 | 309 |
(OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), |
37413 | 310 |
(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
|
311 |
(MalformedInput, "Free Variable"), |
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet
parents:
37413
diff
changeset
|
312 |
(OldSpass, "unrecognized option `-TPTP'"), |
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet
parents:
37413
diff
changeset
|
313 |
(OldSpass, "Unrecognized option TPTP")], |
36382 | 314 |
max_axiom_clauses = 40, |
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
|
315 |
prefers_theory_relevant = true} |
37414
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet
parents:
37413
diff
changeset
|
316 |
val spass = tptp_prover "spass" spass_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
317 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
318 |
(* Vampire *) |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
319 |
|
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
320 |
val vampire_config : prover_config = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
321 |
{home_var = "VAMPIRE_HOME", |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
322 |
executable = "vampire", |
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
|
323 |
arguments = fn _ => fn timeout => |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
324 |
"--output_syntax tptp --mode casc -t " ^ |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
325 |
string_of_int (to_generous_secs timeout), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
326 |
proof_delims = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
327 |
[("=========== Refutation ==========", |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
328 |
"======= End of refutation ======="), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
329 |
("% SZS output start Refutation", "% SZS output end Refutation")], |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
330 |
known_failures = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
331 |
[(Unprovable, "UNPROVABLE"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
332 |
(IncompleteUnprovable, "CANNOT PROVE"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
333 |
(Unprovable, "Satisfiability detected"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
334 |
(OutOfResources, "Refutation not found")], |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
335 |
max_axiom_clauses = 60, |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
336 |
prefers_theory_relevant = false} |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
337 |
val vampire = tptp_prover "vampire" vampire_config |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
338 |
|
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
339 |
(* Remote prover invocation via SystemOnTPTP *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
340 |
|
36376 | 341 |
val systems = Synchronized.var "atp_systems" ([]: string list); |
31835 | 342 |
|
343 |
fun get_systems () = |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
344 |
case bash_output "\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w" of |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
345 |
(answer, 0) => split_lines answer |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
346 |
| (answer, _) => |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
347 |
error ("Failed to get available systems at SystemOnTPTP:\n" ^ answer) |
31835 | 348 |
|
35867 | 349 |
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
|
350 |
Synchronized.change systems (fn _ => get_systems ()) |
31835 | 351 |
|
352 |
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
|
353 |
(if null systems then get_systems () else systems) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
354 |
|> `(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
|
355 |
|
32948 | 356 |
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
|
357 |
(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
|
358 |
NONE => error ("System " ^ quote prefix ^ " not available at SystemOnTPTP.") |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
359 |
| SOME sys => sys); |
31835 | 360 |
|
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
361 |
val remote_known_failures = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
362 |
[(TimedOut, "says Timeout"), |
36377
b3dce4c715d0
now rename the file "atp_wrapper.ML" to "atp_systems.ML" + fix typo in "SystemOnTPTP" script
blanchet
parents:
36376
diff
changeset
|
363 |
(MalformedOutput, "Remote script could not extract proof")] |
35865 | 364 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
365 |
fun remote_config atp_prefix args |
36382 | 366 |
({proof_delims, known_failures, max_axiom_clauses, |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
367 |
prefers_theory_relevant, ...} : prover_config) : prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
368 |
{home_var = "ISABELLE_ATP_MANAGER", |
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
|
369 |
executable = "SystemOnTPTP", |
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
|
370 |
arguments = fn _ => fn timeout => |
36382 | 371 |
args ^ " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^ |
372 |
the_system atp_prefix, |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
373 |
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
|
374 |
known_failures = remote_known_failures @ known_failures, |
36382 | 375 |
max_axiom_clauses = max_axiom_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
|
376 |
prefers_theory_relevant = prefers_theory_relevant} |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
377 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
378 |
fun remote_tptp_prover prover atp_prefix args config = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
379 |
tptp_prover (remotify (fst prover)) (remote_config atp_prefix args config) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
380 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
381 |
val remote_e = remote_tptp_prover e "EP---" "" e_config |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
382 |
val remote_spass = remote_tptp_prover spass "SPASS---" "-x" spass_config |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
383 |
val remote_vampire = remote_tptp_prover vampire "Vampire---9" "" vampire_config |
28592 | 384 |
|
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
385 |
fun maybe_remote (name, _) ({home_var, ...} : prover_config) = |
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
386 |
name |> getenv home_var = "" ? remotify |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
387 |
|
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
388 |
fun default_atps_param_value () = |
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
389 |
space_implode " " [maybe_remote e e_config, maybe_remote spass spass_config, |
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
390 |
remotify (fst vampire)] |
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
391 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
392 |
val provers = [e, spass, vampire, remote_e, remote_spass, remote_vampire] |
35867 | 393 |
val prover_setup = fold add_prover provers |
394 |
||
395 |
val setup = |
|
36376 | 396 |
dest_dir_setup |
35867 | 397 |
#> problem_prefix_setup |
398 |
#> measure_runtime_setup |
|
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
399 |
#> prover_setup |
35867 | 400 |
|
28592 | 401 |
end; |