author | blanchet |
Mon, 26 Jul 2010 20:07:31 +0200 | |
changeset 38001 | a9b47b85ca24 |
parent 38000 | c0b9efa8bfca |
child 38002 | 31705eccee23 |
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 |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
10 |
val trace : bool Unsynchronized.ref |
36376 | 11 |
val dest_dir : string Config.T |
35867 | 12 |
val problem_prefix : string Config.T |
13 |
val measure_runtime : bool Config.T |
|
14 |
val refresh_systems_on_tptp : unit -> unit |
|
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
15 |
val default_atps_param_value : unit -> string |
35867 | 16 |
val setup : theory -> theory |
28592 | 17 |
end; |
18 |
||
36376 | 19 |
structure ATP_Systems : ATP_SYSTEMS = |
28592 | 20 |
struct |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
21 |
|
37578
9367cb36b1c4
renamed "Sledgehammer_FOL_Clauses" to "Metis_Clauses", so that Metis doesn't depend on Sledgehammer
blanchet
parents:
37577
diff
changeset
|
22 |
open Metis_Clauses |
36187 | 23 |
open Sledgehammer_Util |
35865 | 24 |
open Sledgehammer_Fact_Filter |
37578
9367cb36b1c4
renamed "Sledgehammer_FOL_Clauses" to "Metis_Clauses", so that Metis doesn't depend on Sledgehammer
blanchet
parents:
37577
diff
changeset
|
25 |
open Sledgehammer_TPTP_Format |
35865 | 26 |
open Sledgehammer_Proof_Reconstruct |
35867 | 27 |
open ATP_Manager |
35826 | 28 |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
29 |
val trace = Unsynchronized.ref false |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
30 |
fun trace_msg msg = if !trace then tracing (msg ()) else () |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
31 |
|
36376 | 32 |
(** generic ATP **) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
33 |
|
32944
ecc0705174c2
clarified File.platform_path vs. File.shell_path;
wenzelm
parents:
32942
diff
changeset
|
34 |
(* external problem files *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
35 |
|
36376 | 36 |
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
|
37 |
(*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
|
38 |
|
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
39 |
val (problem_prefix, problem_prefix_setup) = |
36001 | 40 |
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
|
41 |
|
33247 | 42 |
val (measure_runtime, measure_runtime_setup) = |
36001 | 43 |
Attrib.config_bool "atp_measure_runtime" (K false); |
33247 | 44 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
45 |
|
35867 | 46 |
(* prover configuration *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
47 |
|
32941
72d48e333b77
eliminated extraneous wrapping of public records;
wenzelm
parents:
32936
diff
changeset
|
48 |
type prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
49 |
{home_var: string, |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
50 |
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
|
51 |
arguments: bool -> Time.time -> string, |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
52 |
proof_delims: (string * string) list, |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
53 |
known_failures: (failure * string) list, |
36382 | 54 |
max_axiom_clauses: int, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
55 |
prefers_theory_relevant: bool, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
56 |
explicit_forall: 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." |
|
37627 | 98 |
| string_for_failure CantConnect = "Can't connect to remote ATP." |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
99 |
| string_for_failure TimedOut = "Timed out." |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
100 |
| 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
|
101 |
| string_for_failure OldSpass = |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36382
diff
changeset
|
102 |
(* FIXME: Change the error message below to point to the Isabelle download |
37997 | 103 |
page once the package is there. *) |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
104 |
"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
|
105 |
\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
|
106 |
\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
|
107 |
\\"spass-3.7\" directory's full path to \"" ^ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
108 |
Path.implode (Path.expand (Path.appends |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
109 |
(Path.variable "ISABELLE_HOME_USER" :: |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
110 |
map Path.basic ["etc", "components"]))) ^ |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
111 |
"\" on a line of its own." |
36965 | 112 |
| string_for_failure MalformedInput = |
113 |
"Internal Sledgehammer error: The ATP problem is malformed. Please report \ |
|
114 |
\this to the Isabelle developers." |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
115 |
| 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
|
116 |
| string_for_failure UnknownError = "Error: An unknown ATP error occurred." |
35865 | 117 |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
118 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
119 |
(* Clause preparation *) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
120 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
121 |
fun make_clause_table xs = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
122 |
fold (Termtab.update o `(prop_of o snd)) xs Termtab.empty |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
123 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
124 |
(* Remove existing axiom clauses from the conjecture clauses, as this can |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
125 |
dramatically boost an ATP's performance (for some reason). *) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
126 |
fun subtract_cls ax_clauses = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
127 |
filter_out (Termtab.defined (make_clause_table ax_clauses) o prop_of) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
128 |
|
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
129 |
fun combformula_for_prop thy = |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
130 |
let |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
131 |
val do_term = combterm_from_term thy |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
132 |
fun do_quant bs q s T t' = |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
133 |
do_formula ((s, T) :: bs) t' |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
134 |
#>> (fn phi => AQuant (q, [`make_bound_var s], phi)) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
135 |
and do_conn bs c t1 t2 = |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
136 |
do_formula bs t1 ##>> do_formula bs t2 |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
137 |
#>> (fn (phi1, phi2) => AConn (c, [phi1, phi2])) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
138 |
and do_formula bs t = |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
139 |
case t of |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
140 |
@{const Not} $ t1 => |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
141 |
do_formula bs t1 #>> (fn phi => AConn (ANot, [phi])) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
142 |
| Const (@{const_name All}, _) $ Abs (s, T, t') => |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
143 |
do_quant bs AForall s T t' |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
144 |
| Const (@{const_name Ex}, _) $ Abs (s, T, t') => |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
145 |
do_quant bs AExists s T t' |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
146 |
| @{const "op &"} $ t1 $ t2 => do_conn bs AAnd t1 t2 |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
147 |
| @{const "op |"} $ t1 $ t2 => do_conn bs AOr t1 t2 |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
148 |
| @{const "op -->"} $ t1 $ t2 => do_conn bs AImplies t1 t2 |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
149 |
| Const (@{const_name "op ="}, Type (_, [@{typ bool}, _])) $ t1 $ t2 => |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
150 |
do_conn bs AIff t1 t2 |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
151 |
| _ => (fn ts => do_term bs (Envir.eta_contract t) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
152 |
|>> APred ||> union (op =) ts) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
153 |
in do_formula [] end |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
154 |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
155 |
(* Converts an elim-rule into an equivalent theorem that does not have the |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
156 |
predicate variable. Leaves other theorems unchanged. We simply instantiate |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
157 |
the conclusion variable to False. (Cf. "transform_elim_term" in |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
158 |
"ATP_Systems".) *) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
159 |
(* FIXME: test! *) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
160 |
fun transform_elim_term t = |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
161 |
case Logic.strip_imp_concl t of |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
162 |
@{const Trueprop} $ Var (z, @{typ bool}) => |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
163 |
subst_Vars [(z, @{const True})] t |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
164 |
| Var (z, @{typ prop}) => subst_Vars [(z, @{prop True})] t |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
165 |
| _ => t |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
166 |
|
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
167 |
(* Removes the lambdas from an equation of the form "t = (%x. u)". |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
168 |
(Cf. "extensionalize_theorem" in "Clausifier".) *) |
38000
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
169 |
fun extensionalize_term t = |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
170 |
let |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
171 |
fun aux j (Const (@{const_name "op ="}, Type (_, [Type (_, [_, T']), _])) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
172 |
$ t2 $ Abs (s, var_T, t')) = |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
173 |
let val var_t = Var (("x", j), var_T) in |
38000
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
174 |
Const (@{const_name "op ="}, T' --> T' --> HOLogic.boolT) |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
175 |
$ betapply (t2, var_t) $ subst_bound (var_t, t') |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
176 |
|> aux (j + 1) |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
177 |
end |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
178 |
| aux _ t = t |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
179 |
in aux (maxidx_of_term t + 1) t end |
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
180 |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
181 |
(* FIXME: Guarantee freshness *) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
182 |
fun concealed_bound_name j = "Sledgehammer" ^ Int.toString j |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
183 |
fun conceal_bounds Ts t = |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
184 |
subst_bounds (map (Free o apfst concealed_bound_name) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
185 |
(length Ts - 1 downto 0 ~~ rev Ts), t) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
186 |
fun reveal_bounds Ts = |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
187 |
subst_atomic (map (fn (j, T) => (Free (concealed_bound_name j, T), Bound j)) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
188 |
(0 upto length Ts - 1 ~~ Ts)) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
189 |
|
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
190 |
fun introduce_combinators_in_term thy t = |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
191 |
let |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
192 |
fun aux Ts t = |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
193 |
case t of |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
194 |
@{const Not} $ t1 => @{const Not} $ aux Ts t1 |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
195 |
| (t0 as Const (@{const_name All}, _)) $ Abs (s, T, t') => |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
196 |
t0 $ Abs (s, T, aux (T :: Ts) t') |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
197 |
| (t0 as Const (@{const_name Ex}, _)) $ Abs (s, T, t') => |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
198 |
t0 $ Abs (s, T, aux (T :: Ts) t') |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
199 |
| (t0 as @{const "op &"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
200 |
| (t0 as @{const "op |"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
201 |
| (t0 as @{const "op -->"}) $ t1 $ t2 => t0 $ aux Ts t1 $ aux Ts t2 |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
202 |
| (t0 as Const (@{const_name "op ="}, Type (_, [@{typ bool}, _]))) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
203 |
$ t1 $ t2 => |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
204 |
t0 $ aux Ts t1 $ aux Ts t2 |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
205 |
| _ => if not (exists_subterm (fn Abs _ => true | _ => false) t) then |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
206 |
t |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
207 |
else |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
208 |
t |> conceal_bounds Ts |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
209 |
|> Envir.eta_contract |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
210 |
|> cterm_of thy |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
211 |
|> Clausifier.introduce_combinators_in_cterm |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
212 |
|> prop_of |> Logic.dest_equals |> snd |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
213 |
|> reveal_bounds Ts |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
214 |
handle THM (msg, _, _) => |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
215 |
(* A type variable of sort "{}" will make abstraction |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
216 |
fail. *) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
217 |
t |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
218 |
in t |> not (Meson.is_fol_term thy t) ? aux [] end |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
219 |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
220 |
(* making axiom and conjecture clauses *) |
37997 | 221 |
fun make_clause thy (formula_id, formula_name, kind, t) = |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
222 |
let |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
223 |
(* ### FIXME: introduce combinators and perform other transformations |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
224 |
previously done by Clausifier.to_nnf *) |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
225 |
val t = t |> transform_elim_term |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
226 |
|> Object_Logic.atomize_term thy |
38000
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37999
diff
changeset
|
227 |
|> extensionalize_term |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
228 |
|> introduce_combinators_in_term thy |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
229 |
val (combformula, ctypes_sorts) = combformula_for_prop thy t [] |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
230 |
in |
37997 | 231 |
FOLFormula {formula_name = formula_name, formula_id = formula_id, |
232 |
combformula = combformula, kind = kind, |
|
233 |
ctypes_sorts = ctypes_sorts} |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
234 |
end |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
235 |
|
37997 | 236 |
fun add_axiom_clause thy ((name, k), th) = |
237 |
cons (name, make_clause thy (k, name, Axiom, prop_of th)) |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
238 |
|
37997 | 239 |
fun make_axiom_clauses thy clauses = fold_rev (add_axiom_clause thy) clauses [] |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
240 |
|
37999 | 241 |
fun make_conjecture_clauses thy ts = |
242 |
map2 (fn j => fn t => make_clause thy (j, "conjecture", Conjecture, t)) |
|
243 |
(0 upto length ts - 1) ts |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
244 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
245 |
(** Helper clauses **) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
246 |
|
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
247 |
fun count_combterm (CombConst ((s, _), _, _)) = |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
248 |
Symtab.map_entry s (Integer.add 1) |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
249 |
| count_combterm (CombVar _) = I |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
250 |
| count_combterm (CombApp (t1, t2)) = fold count_combterm [t1, t2] |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
251 |
fun count_combformula (AQuant (_, _, phi)) = count_combformula phi |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
252 |
| count_combformula (AConn (_, phis)) = fold count_combformula phis |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
253 |
| count_combformula (APred tm) = count_combterm tm |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
254 |
fun count_fol_formula (FOLFormula {combformula, ...}) = |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
255 |
count_combformula combformula |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
256 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
257 |
fun cnf_helper_thms thy raw = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
258 |
map (`Thm.get_name_hint) |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37925
diff
changeset
|
259 |
#> (if raw then map (apfst (rpair 0)) else Clausifier.cnf_rules_pairs thy true) |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
260 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
261 |
val optional_helpers = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
262 |
[(["c_COMBI", "c_COMBK"], (false, @{thms COMBI_def COMBK_def})), |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
263 |
(["c_COMBB", "c_COMBC"], (false, @{thms COMBB_def COMBC_def})), |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
264 |
(["c_COMBS"], (false, @{thms COMBS_def}))] |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
265 |
val optional_typed_helpers = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
266 |
[(["c_True", "c_False"], (true, @{thms True_or_False})), |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
267 |
(["c_If"], (true, @{thms if_True if_False True_or_False}))] |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
268 |
val mandatory_helpers = @{thms fequal_imp_equal equal_imp_fequal} |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
269 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
270 |
val init_counters = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
271 |
Symtab.make (maps (maps (map (rpair 0) o fst)) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
272 |
[optional_helpers, optional_typed_helpers]) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
273 |
|
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
274 |
fun get_helper_clauses thy is_FO full_types conjectures axcls = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
275 |
let |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
276 |
val axclauses = map snd (make_axiom_clauses thy axcls) |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
277 |
val ct = fold (fold count_fol_formula) [conjectures, axclauses] |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
278 |
init_counters |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
279 |
fun is_needed c = the (Symtab.lookup ct c) > 0 |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
280 |
val cnfs = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
281 |
(optional_helpers |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
282 |
|> full_types ? append optional_typed_helpers |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
283 |
|> maps (fn (ss, (raw, ths)) => |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
284 |
if exists is_needed ss then cnf_helper_thms thy raw ths |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
285 |
else [])) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
286 |
@ (if is_FO then [] else cnf_helper_thms thy false mandatory_helpers) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
287 |
in map snd (make_axiom_clauses thy cnfs) end |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
288 |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
289 |
fun negate_term (@{const Not} $ t) = t |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
290 |
| negate_term t = @{const Not} $ t |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
291 |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
292 |
(* prepare for passing to writer, |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
293 |
create additional clauses based on the information from extra_cls *) |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
294 |
fun prepare_clauses full_types hyp_ts concl_t axcls extra_cls thy = |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
295 |
let |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
296 |
val goal_t = Logic.list_implies (hyp_ts, concl_t) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
297 |
val is_FO = Meson.is_fol_term thy goal_t |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
298 |
val _ = trace_msg (fn _ => Syntax.string_of_term_global thy goal_t) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
299 |
val axtms = map (prop_of o snd) extra_cls |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
300 |
val subs = tfree_classes_of_terms [goal_t] |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
301 |
val supers = tvar_classes_of_terms axtms |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
302 |
val tycons = type_consts_of_terms thy (goal_t :: axtms) |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
303 |
(*TFrees in conjecture clauses; TVars in axiom clauses*) |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
304 |
val conjectures = |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
305 |
make_conjecture_clauses thy (map negate_term hyp_ts @ [concl_t]) |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
306 |
val (_, extra_clauses) = ListPair.unzip (make_axiom_clauses thy extra_cls) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
307 |
val (clnames, axiom_clauses) = ListPair.unzip (make_axiom_clauses thy axcls) |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
308 |
val helper_clauses = |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
309 |
get_helper_clauses thy is_FO full_types conjectures extra_cls |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
310 |
val (supers', arity_clauses) = make_arity_clauses thy tycons supers |
37925 | 311 |
val class_rel_clauses = make_class_rel_clauses thy subs supers' |
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
312 |
in |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
313 |
(Vector.fromList clnames, |
37925 | 314 |
(conjectures, axiom_clauses, extra_clauses, helper_clauses, |
315 |
class_rel_clauses, arity_clauses)) |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
316 |
end |
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
317 |
|
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
318 |
fun extract_clause_sequence output = |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
319 |
let |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
320 |
val tokens_of = String.tokens (not o Char.isAlphaNum) |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
321 |
fun extract_num ("clause" :: (ss as _ :: _)) = |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
322 |
Int.fromString (List.last ss) |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
323 |
| extract_num _ = NONE |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
324 |
in output |> split_lines |> map_filter (extract_num o tokens_of) end |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
325 |
|
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
326 |
val set_ClauseFormulaRelationN = "set_ClauseFormulaRelation" |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
327 |
|
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
328 |
val parse_clause_formula_pair = |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
329 |
$$ "(" |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ ")" |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
330 |
--| Scan.option ($$ ",") |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
331 |
val parse_clause_formula_relation = |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
332 |
Scan.this_string set_ClauseFormulaRelationN |-- $$ "(" |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
333 |
|-- Scan.repeat parse_clause_formula_pair |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
334 |
val extract_clause_formula_relation = |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
335 |
Substring.full |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
336 |
#> Substring.position set_ClauseFormulaRelationN |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
337 |
#> snd #> Substring.string #> strip_spaces #> explode |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
338 |
#> parse_clause_formula_relation #> fst |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
339 |
|
37989
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
340 |
fun repair_conjecture_shape_and_theorem_names output conjecture_shape |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
341 |
thm_names = |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
342 |
if String.isSubstring set_ClauseFormulaRelationN output then |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
343 |
let |
37996
11c076ea92e9
simplify "conjecture_shape" business, as a result of using FOF instead of CNF
blanchet
parents:
37995
diff
changeset
|
344 |
val j0 = hd conjecture_shape |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
345 |
val seq = extract_clause_sequence output |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
346 |
val name_map = extract_clause_formula_relation output |
37989
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
347 |
fun renumber_conjecture j = |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
348 |
AList.find (op =) name_map (conjecture_prefix ^ Int.toString (j - j0)) |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
349 |
|> the_single |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
350 |
|> (fn s => find_index (curry (op =) s) seq + 1) |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
351 |
in |
37996
11c076ea92e9
simplify "conjecture_shape" business, as a result of using FOF instead of CNF
blanchet
parents:
37995
diff
changeset
|
352 |
(conjecture_shape |> map renumber_conjecture, |
37989
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
353 |
seq |> map (the o AList.lookup (op =) name_map) |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
354 |
|> map (fn s => case try (unprefix axiom_prefix) s of |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
355 |
SOME s' => undo_ascii_of s' |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
356 |
| NONE => "") |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
357 |
|> Vector.fromList) |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
358 |
end |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
359 |
else |
37989
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
360 |
(conjecture_shape, thm_names) |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
361 |
|
37623
295f3a9b44b6
move functions not needed by Metis out of "Metis_Clauses"
blanchet
parents:
37621
diff
changeset
|
362 |
|
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
363 |
(* generic TPTP-based provers *) |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
364 |
|
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
365 |
fun generic_tptp_prover |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
366 |
(name, {home_var, executable, arguments, proof_delims, known_failures, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
367 |
max_axiom_clauses, prefers_theory_relevant, explicit_forall}) |
37580
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
368 |
({debug, overlord, full_types, explicit_apply, relevance_threshold, |
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
369 |
relevance_convergence, theory_relevant, defs_relevant, isar_proof, |
c2c1caff5dea
got rid of "respect_no_atp" option, which even I don't use
blanchet
parents:
37578
diff
changeset
|
370 |
isar_shrink_factor, ...} : params) |
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
371 |
minimize_command timeout |
35969 | 372 |
({subgoal, goal, relevance_override, axiom_clauses, filtered_clauses} |
373 |
: problem) = |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
374 |
let |
31750 | 375 |
(* get clauses and prepare them for writing *) |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37480
diff
changeset
|
376 |
val (ctxt, (_, th)) = goal; |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
377 |
val thy = ProofContext.theory_of ctxt; |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
378 |
(* ### FIXME: (1) preprocessing for "if" etc. *) |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
379 |
val (params, hyp_ts, concl_t) = strip_subgoal th subgoal |
31752
19a5f1c8a844
use results of relevance-filter to determine additional clauses;
immler@in.tum.de
parents:
31751
diff
changeset
|
380 |
val the_filtered_clauses = |
37621 | 381 |
case filtered_clauses of |
382 |
SOME fcls => fcls |
|
383 |
| NONE => relevant_facts full_types relevance_threshold |
|
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
384 |
relevance_convergence defs_relevant max_axiom_clauses |
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
385 |
(the_default prefers_theory_relevant theory_relevant) |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
386 |
relevance_override goal hyp_ts concl_t |
37926
e6ff246c0cdb
renamings + only need second component of name pool to reconstruct proofs
blanchet
parents:
37925
diff
changeset
|
387 |
|> Clausifier.cnf_rules_pairs thy true |
37506
32a1ee39c49b
missing "Unsynchronized" + make exception take a unit
blanchet
parents:
37499
diff
changeset
|
388 |
val the_axiom_clauses = axiom_clauses |> the_default the_filtered_clauses |
35969 | 389 |
val (internal_thm_names, clauses) = |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
390 |
prepare_clauses full_types hyp_ts concl_t the_axiom_clauses |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
391 |
the_filtered_clauses thy |
31750 | 392 |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
393 |
(* path to unique problem file *) |
36376 | 394 |
val the_dest_dir = if overlord then getenv "ISABELLE_HOME_USER" |
395 |
else Config.get ctxt dest_dir; |
|
396 |
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
|
397 |
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
|
398 |
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
|
399 |
val probfile = |
36568
d495d2e1f0a6
in "overlord" mode: ignore problem prefix specified in the .thy file
blanchet
parents:
36552
diff
changeset
|
400 |
Path.basic ((if overlord then "prob_" ^ name |
d495d2e1f0a6
in "overlord" mode: ignore problem prefix specified in the .thy file
blanchet
parents:
36552
diff
changeset
|
401 |
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
|
402 |
^ "_" ^ string_of_int nr) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
403 |
in |
36376 | 404 |
if the_dest_dir = "" then File.tmp_path probfile |
405 |
else if File.exists (Path.explode the_dest_dir) |
|
406 |
then Path.append (Path.explode the_dest_dir) probfile |
|
407 |
else error ("No such directory: " ^ the_dest_dir ^ ".") |
|
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
408 |
end; |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
409 |
|
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
410 |
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
|
411 |
val command = Path.explode (home ^ "/" ^ executable) |
31750 | 412 |
(* 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
|
413 |
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
|
414 |
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
|
415 |
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
|
416 |
" " ^ 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
|
417 |
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
|
418 |
(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
|
419 |
"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
|
420 |
else |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
421 |
"exec " ^ core) ^ " 2>&1" |
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
|
422 |
end |
32510 | 423 |
fun split_time s = |
424 |
let |
|
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
425 |
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
|
426 |
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
|
427 |
fun as_num f = f >> (fst o read_int); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
428 |
val num = as_num (Scan.many1 Symbol.is_ascii_digit); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
429 |
val digit = Scan.one Symbol.is_ascii_digit; |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
430 |
val num3 = as_num (digit ::: digit ::: (digit >> single)); |
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
431 |
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
|
432 |
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
|
433 |
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
|
434 |
fun run_on probfile = |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
435 |
if home = "" then |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
436 |
error ("The environment variable " ^ quote home_var ^ " is not set.") |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
437 |
else 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
|
438 |
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
|
439 |
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
|
440 |
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
|
441 |
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
|
442 |
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
|
443 |
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
|
444 |
|>> (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
|
445 |
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
|
446 |
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
|
447 |
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
|
448 |
|>> (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
|
449 |
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
|
450 |
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
|
451 |
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
|
452 |
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
|
453 |
in (output, msecs, proof, outcome) end |
37643
f576af716aa6
rewrote the TPTP problem generation code more or less from scratch;
blanchet
parents:
37628
diff
changeset
|
454 |
val readable_names = debug andalso overlord |
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
|
455 |
val (pool, conjecture_offset) = |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
456 |
write_tptp_file thy readable_names explicit_forall full_types |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
457 |
explicit_apply probfile clauses |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
458 |
val conjecture_shape = |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
459 |
conjecture_offset + 1 upto conjecture_offset + length hyp_ts + 1 |
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
|
460 |
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
|
461 |
do_run false |
37550
fc2f979b9a08
split SPASS time slot between SOS and non-SOS, in case SOS times out
blanchet
parents:
37514
diff
changeset
|
462 |
|> (fn (_, msecs0, _, SOME _) => |
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
|
463 |
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
|
464 |
|> (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
|
465 |
(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
|
466 |
| 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
|
467 |
in ((pool, conjecture_shape), result) end |
36382 | 468 |
else |
469 |
error ("Bad executable: " ^ Path.implode command ^ "."); |
|
28592 | 470 |
|
36167
c1a35be8e476
make Sledgehammer's output more debugging friendly
blanchet
parents:
36143
diff
changeset
|
471 |
(* 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
|
472 |
the proof file too. *) |
36376 | 473 |
fun cleanup probfile = |
474 |
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
|
475 |
fun export probfile (_, (output, _, _, _)) = |
36376 | 476 |
if the_dest_dir = "" then |
36187 | 477 |
() |
478 |
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
|
479 |
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
|
480 |
|
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
|
481 |
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
|
482 |
with_path cleanup export run_on (prob_pathname subgoal) |
37989
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
483 |
val (conjecture_shape, internal_thm_names) = |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
484 |
repair_conjecture_shape_and_theorem_names output conjecture_shape |
ca3041b0f445
reorder SPASS conjectures correctly, based on Flotter output
blanchet
parents:
37962
diff
changeset
|
485 |
internal_thm_names |
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
|
486 |
|
35969 | 487 |
val (message, relevant_thm_names) = |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
488 |
case outcome of |
36400
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
489 |
NONE => |
c5bae529f967
rename options and keep track of conjecture shape (to facilitate proof reconstruction)
blanchet
parents:
36393
diff
changeset
|
490 |
proof_text isar_proof |
37480 | 491 |
(pool, debug, isar_shrink_factor, ctxt, conjecture_shape) |
492 |
(full_types, minimize_command, proof, internal_thm_names, th, |
|
493 |
subgoal) |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
494 |
| 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
|
495 |
in |
36393
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
blanchet
parents:
36382
diff
changeset
|
496 |
{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
|
497 |
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
|
498 |
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
|
499 |
conjecture_shape = conjecture_shape, |
35969 | 500 |
filtered_clauses = the_filtered_clauses} |
37499
5ff37037fbec
merge "generic_prover" and "generic_tptp_prover"
blanchet
parents:
37498
diff
changeset
|
501 |
end |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
502 |
|
35969 | 503 |
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
|
504 |
|
36382 | 505 |
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
|
506 |
|
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
507 |
(* E prover *) |
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
508 |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
509 |
val tstp_proof_delims = |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
510 |
("# SZS output start CNFRefutation.", "# SZS output end CNFRefutation") |
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
511 |
|
35969 | 512 |
val e_config : prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
513 |
{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
|
514 |
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
|
515 |
arguments = fn _ => fn timeout => |
36382 | 516 |
"--tstp-in --tstp-out -l5 -xAutoDev -tAutoDev --silent --cpu-limit=" ^ |
517 |
string_of_int (to_generous_secs timeout), |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
518 |
proof_delims = [tstp_proof_delims], |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
519 |
known_failures = |
37995
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
520 |
[(Unprovable, "SZS status: CounterSatisfiable"), |
06f02b15ef8a
generate full first-order formulas (FOF) in Sledgehammer
blanchet
parents:
37994
diff
changeset
|
521 |
(Unprovable, "SZS status CounterSatisfiable"), |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
522 |
(TimedOut, "Failure: Resource limit exceeded (time)"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
523 |
(TimedOut, "time limit exceeded"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
524 |
(OutOfResources, |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
525 |
"# Cannot determine problem status within resource limit"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
526 |
(OutOfResources, "SZS status: ResourceOut"), |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
527 |
(OutOfResources, "SZS status ResourceOut")], |
36382 | 528 |
max_axiom_clauses = 100, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
529 |
prefers_theory_relevant = false, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
530 |
explicit_forall = false} |
35969 | 531 |
val e = tptp_prover "e" e_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
532 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
533 |
|
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
|
534 |
(* 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
|
535 |
counteracting the presence of "hAPP". *) |
37498
b426cbdb5a23
removed Sledgehammer's support for the DFG syntax;
blanchet
parents:
37480
diff
changeset
|
536 |
val spass_config : prover_config = |
37962
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
537 |
{home_var = "ISABELLE_ATP_MANAGER", |
d7dbe01f48d7
keep track of clause numbers for SPASS now that we generate FOF rather than CNF problems;
blanchet
parents:
37926
diff
changeset
|
538 |
executable = "SPASS_TPTP", |
37550
fc2f979b9a08
split SPASS time slot between SOS and non-SOS, in case SOS times out
blanchet
parents:
37514
diff
changeset
|
539 |
(* "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
|
540 |
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
|
541 |
("-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
|
542 |
\-VarWeight=3 -TimeLimit=" ^ |
fc2f979b9a08
split SPASS time slot between SOS and non-SOS, in case SOS times out
blanchet
parents:
37514
diff
changeset
|
543 |
string_of_int (to_generous_secs timeout 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
|
544 |
|> not complete ? prefix "-SOS=1 ", |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
545 |
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
|
546 |
known_failures = |
37413 | 547 |
[(IncompleteUnprovable, "SPASS beiseite: Completion found"), |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
548 |
(TimedOut, "SPASS beiseite: Ran out of time"), |
36965 | 549 |
(OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), |
37413 | 550 |
(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
|
551 |
(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
|
552 |
(OldSpass, "tptp2dfg")], |
36382 | 553 |
max_axiom_clauses = 40, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
554 |
prefers_theory_relevant = true, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
555 |
explicit_forall = true} |
37414
d0cea0796295
expect SPASS 3.7, and give a friendly warning if an older version is used
blanchet
parents:
37413
diff
changeset
|
556 |
val spass = tptp_prover "spass" spass_config |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
557 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
558 |
(* Vampire *) |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
559 |
|
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
560 |
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
|
561 |
{home_var = "VAMPIRE_HOME", |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
562 |
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
|
563 |
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
|
564 |
"--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
|
565 |
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
|
566 |
proof_delims = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
567 |
[("=========== Refutation ==========", |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
568 |
"======= End of refutation ======="), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
569 |
("% 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
|
570 |
known_failures = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
571 |
[(Unprovable, "UNPROVABLE"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
572 |
(IncompleteUnprovable, "CANNOT PROVE"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
573 |
(Unprovable, "Satisfiability detected"), |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
574 |
(OutOfResources, "Refutation not found")], |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
575 |
max_axiom_clauses = 60, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
576 |
prefers_theory_relevant = false, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
577 |
explicit_forall = false} |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
578 |
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
|
579 |
|
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
580 |
(* Remote prover invocation via SystemOnTPTP *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
581 |
|
36376 | 582 |
val systems = Synchronized.var "atp_systems" ([]: string list); |
31835 | 583 |
|
584 |
fun get_systems () = |
|
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
585 |
case bash_output "\"$ISABELLE_ATP_MANAGER/SystemOnTPTP\" -w" of |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
586 |
(answer, 0) => split_lines answer |
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
587 |
| (answer, _) => |
37627 | 588 |
error ("Failed to get available systems at SystemOnTPTP:\n" ^ |
589 |
perhaps (try (unsuffix "\n")) answer) |
|
31835 | 590 |
|
35867 | 591 |
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
|
592 |
Synchronized.change systems (fn _ => get_systems ()) |
31835 | 593 |
|
594 |
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
|
595 |
(if null systems then get_systems () else systems) |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
596 |
|> `(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
|
597 |
|
32948 | 598 |
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
|
599 |
(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
|
600 |
NONE => error ("System " ^ quote prefix ^ " not available at SystemOnTPTP.") |
32942
b6711ec9de26
misc tuning and recovery of Isabelle coding style;
wenzelm
parents:
32941
diff
changeset
|
601 |
| SOME sys => sys); |
31835 | 602 |
|
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
603 |
val remote_known_failures = |
37627 | 604 |
[(CantConnect, "HTTP-Error"), |
605 |
(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
|
606 |
(MalformedOutput, "Remote script could not extract proof")] |
35865 | 607 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
608 |
fun remote_config atp_prefix args |
36382 | 609 |
({proof_delims, known_failures, max_axiom_clauses, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
610 |
prefers_theory_relevant, explicit_forall, ...} : prover_config) |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
611 |
: prover_config = |
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
612 |
{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
|
613 |
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
|
614 |
arguments = fn _ => fn timeout => |
36382 | 615 |
args ^ " -t " ^ string_of_int (to_generous_secs timeout) ^ " -s " ^ |
616 |
the_system atp_prefix, |
|
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
617 |
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
|
618 |
known_failures = remote_known_failures @ known_failures, |
36382 | 619 |
max_axiom_clauses = max_axiom_clauses, |
37994
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
620 |
prefers_theory_relevant = prefers_theory_relevant, |
b04307085a09
make TPTP generator accept full first-order formulas
blanchet
parents:
37989
diff
changeset
|
621 |
explicit_forall = explicit_forall} |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
622 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
623 |
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
|
624 |
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
|
625 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
626 |
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
|
627 |
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
|
628 |
val remote_vampire = remote_tptp_prover vampire "Vampire---9" "" vampire_config |
28592 | 629 |
|
36917
8674cdb0b8cc
query _HOME environment variables at run-time, not at build-time
blanchet
parents:
36910
diff
changeset
|
630 |
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
|
631 |
name |> getenv home_var = "" ? remotify |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
632 |
|
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
633 |
fun default_atps_param_value () = |
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
634 |
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
|
635 |
remotify (fst vampire)] |
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
636 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
637 |
val provers = [e, spass, vampire, remote_e, remote_spass, remote_vampire] |
35867 | 638 |
val prover_setup = fold add_prover provers |
639 |
||
640 |
val setup = |
|
36376 | 641 |
dest_dir_setup |
35867 | 642 |
#> problem_prefix_setup |
643 |
#> measure_runtime_setup |
|
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
644 |
#> prover_setup |
35867 | 645 |
|
28592 | 646 |
end; |