author | blanchet |
Sun, 04 May 2014 18:14:59 +0200 | |
changeset 56848 | 67e6803e3765 |
parent 56404 | 9cb137ec6ec8 |
child 57008 | 10f68b83b474 |
permissions | -rw-r--r-- |
38047 | 1 |
(* Title: HOL/Tools/ATP/atp_systems.ML |
28592 | 2 |
Author: Fabian Immler, TU Muenchen |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
3 |
Author: Jasmin Blanchette, TU Muenchen |
28592 | 4 |
|
36376 | 5 |
Setup for supported ATPs. |
28592 | 6 |
*) |
7 |
||
36376 | 8 |
signature ATP_SYSTEMS = |
28592 | 9 |
sig |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
10 |
type term_order = ATP_Problem.term_order |
45301
866b075aa99b
added sorted DFG output for coming version of SPASS
blanchet
parents:
45300
diff
changeset
|
11 |
type atp_format = ATP_Problem.atp_format |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
12 |
type atp_formula_role = ATP_Problem.atp_formula_role |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
13 |
type atp_failure = ATP_Proof.atp_failure |
38023 | 14 |
|
51011 | 15 |
type slice_spec = (int * string) * atp_format * string * string * bool |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
16 |
type atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
17 |
{exec : unit -> string list * string list, |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
18 |
arguments : |
50927 | 19 |
Proof.context -> bool -> string -> Time.time -> string |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
20 |
-> term_order * (unit -> (string * int) list) |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
21 |
* (unit -> (string * real) list) -> string, |
42578
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
blanchet
parents:
42577
diff
changeset
|
22 |
proof_delims : (string * string) list, |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
23 |
known_failures : (atp_failure * string) list, |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
24 |
prem_role : atp_formula_role, |
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48715
diff
changeset
|
25 |
best_slices : Proof.context -> (real * (slice_spec * string)) list, |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
26 |
best_max_mono_iters : int, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
27 |
best_max_new_mono_instances : int} |
38023 | 28 |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
29 |
val default_max_mono_iters : int |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
30 |
val default_max_new_mono_instances : int |
44099 | 31 |
val force_sos : bool Config.T |
47032 | 32 |
val term_order : string Config.T |
43566
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents:
43529
diff
changeset
|
33 |
val e_smartN : string |
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents:
43529
diff
changeset
|
34 |
val e_autoN : string |
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents:
43529
diff
changeset
|
35 |
val e_fun_weightN : string |
a818d5a34cca
filter out some tautologies using an ATP, especially for those theories that are known for producing such things
blanchet
parents:
43529
diff
changeset
|
36 |
val e_sym_offset_weightN : string |
47032 | 37 |
val e_selection_heuristic : string Config.T |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
38 |
val e_default_fun_weight : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
39 |
val e_fun_weight_base : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
40 |
val e_fun_weight_span : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
41 |
val e_default_sym_offs_weight : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
42 |
val e_sym_offs_weight_base : real Config.T |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
43 |
val e_sym_offs_weight_span : real Config.T |
50950 | 44 |
val spass_H1SOS : string |
45 |
val spass_H2 : string |
|
46 |
val spass_H2LR0LT0 : string |
|
47 |
val spass_H2NuVS0 : string |
|
48 |
val spass_H2NuVS0Red2 : string |
|
49 |
val spass_H2SOS : string |
|
50 |
val spass_extra_options : string Config.T |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
51 |
val agsyholN : string |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
52 |
val alt_ergoN : string |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
53 |
val dummy_thfN : string |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
54 |
val eN : string |
48651 | 55 |
val e_malesN : string |
50927 | 56 |
val e_parN : string |
44590 | 57 |
val e_sineN : string |
58 |
val e_tofofN : string |
|
45338 | 59 |
val iproverN : string |
60 |
val iprover_eqN : string |
|
44590 | 61 |
val leo2N : string |
62 |
val satallaxN : string |
|
63 |
val snarkN : string |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
64 |
val spassN : string |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
65 |
val spass_pirateN : string |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
66 |
val vampireN : string |
42938 | 67 |
val waldmeisterN : string |
44423
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents:
44422
diff
changeset
|
68 |
val z3_tptpN : string |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
69 |
val remote_prefix : string |
41738
eb98c60a6cf0
added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
blanchet
parents:
41727
diff
changeset
|
70 |
val remote_atp : |
eb98c60a6cf0
added experimental "remote_z3_atp", Sutcliffe's TPTP-syntax-aware wrapper for Z3 -- allows to do head-to-head comparison of Sledgehammer's ATP translation and of Sascha's SMT translation
blanchet
parents:
41727
diff
changeset
|
71 |
string -> string -> string list -> (string * string) list |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
72 |
-> (atp_failure * string) list -> atp_formula_role |
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
73 |
-> (Proof.context -> slice_spec * string) -> string * (unit -> atp_config) |
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
74 |
val add_atp : string * (unit -> atp_config) -> theory -> theory |
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
75 |
val get_atp : theory -> string -> (unit -> atp_config) |
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41725
diff
changeset
|
76 |
val supported_atps : theory -> string list |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
77 |
val is_atp_installed : theory -> string -> bool |
35867 | 78 |
val refresh_systems_on_tptp : unit -> unit |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
79 |
val effective_term_order : Proof.context -> string -> term_order |
35867 | 80 |
val setup : theory -> theory |
28592 | 81 |
end; |
82 |
||
36376 | 83 |
structure ATP_Systems : ATP_SYSTEMS = |
28592 | 84 |
struct |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
85 |
|
42577
78414ec6fa4e
made the format (TFF or FOF) of the TPTP problem a global argument of the problem again and have the ATPs report which formats they support
blanchet
parents:
42571
diff
changeset
|
86 |
open ATP_Problem |
39491
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents:
39375
diff
changeset
|
87 |
open ATP_Proof |
46320 | 88 |
open ATP_Problem_Generate |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
89 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
90 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
91 |
(* ATP configuration *) |
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
92 |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
93 |
val default_max_mono_iters = 3 (* FUDGE *) |
53515
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents:
53225
diff
changeset
|
94 |
val default_max_new_mono_instances = 100 (* FUDGE *) |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
95 |
|
51011 | 96 |
type slice_spec = (int * string) * atp_format * string * string * bool |
46409
d4754183ccce
made option available to users (mostly for experiments)
blanchet
parents:
46407
diff
changeset
|
97 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
98 |
type atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
99 |
{exec : unit -> string list * string list, |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
100 |
arguments : |
50927 | 101 |
Proof.context -> bool -> string -> Time.time -> string |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
102 |
-> term_order * (unit -> (string * int) list) |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
103 |
* (unit -> (string * real) list) -> string, |
42578
1eaf4d437d4c
define type system in ATP module so that ATPs can suggest a type system
blanchet
parents:
42577
diff
changeset
|
104 |
proof_delims : (string * string) list, |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
105 |
known_failures : (atp_failure * string) list, |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
106 |
prem_role : atp_formula_role, |
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48715
diff
changeset
|
107 |
best_slices : Proof.context -> (real * (slice_spec * string)) list, |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
108 |
best_max_mono_iters : int, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
109 |
best_max_new_mono_instances : int} |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
110 |
|
42723 | 111 |
(* "best_slices" must be found empirically, taking a wholistic approach since |
51011 | 112 |
the ATPs are run in parallel. Each slice has the format |
113 |
||
114 |
(time_frac, ((max_facts, fact_filter), format, type_enc, |
|
115 |
lam_trans, uncurried_aliases), extra) |
|
116 |
||
117 |
where |
|
118 |
||
119 |
time_frac = faction of the time available given to the slice (which should |
|
120 |
add up to 1.0) |
|
121 |
||
122 |
extra = extra information to the prover (e.g., SOS or no SOS). |
|
42723 | 123 |
|
124 |
The last slice should be the most "normal" one, because it will get all the |
|
43569
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents:
43567
diff
changeset
|
125 |
time available if the other slices fail early and also because it is used if |
b342cd125533
removed "full_types" option from Sledgehammer, now that virtually sound encodings are used as the default anyway
blanchet
parents:
43567
diff
changeset
|
126 |
slicing is disabled (e.g., by the minimizer). *) |
42710
84fcce345b5d
further improved type system setup based on Judgment Days
blanchet
parents:
42709
diff
changeset
|
127 |
|
51011 | 128 |
val mepoN = "mepo" |
129 |
val mashN = "mash" |
|
130 |
val meshN = "mesh" |
|
131 |
||
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
132 |
val tstp_proof_delims = |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
133 |
[("% SZS output start CNFRefutation", "% SZS output end CNFRefutation"), |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
134 |
("% SZS output start Refutation", "% SZS output end Refutation"), |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
135 |
("% SZS output start Proof", "% SZS output end Proof")] |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
136 |
|
38061
685d1f0f75b3
handle Perl and "libwww-perl" failures more gracefully, giving the user some clues about what goes on
blanchet
parents:
38049
diff
changeset
|
137 |
val known_perl_failures = |
38094 | 138 |
[(CantConnect, "HTTP error"), |
139 |
(NoPerl, "env: perl"), |
|
38065 | 140 |
(NoLibwwwPerl, "Can't locate HTTP")] |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
141 |
|
45203 | 142 |
fun known_szs_failures wrap = |
143 |
[(Unprovable, wrap "CounterSatisfiable"), |
|
144 |
(Unprovable, wrap "Satisfiable"), |
|
145 |
(GaveUp, wrap "GaveUp"), |
|
146 |
(GaveUp, wrap "Unknown"), |
|
147 |
(GaveUp, wrap "Incomplete"), |
|
148 |
(ProofMissing, wrap "Theorem"), |
|
149 |
(ProofMissing, wrap "Unsatisfiable"), |
|
150 |
(TimedOut, wrap "Timeout"), |
|
151 |
(Inappropriate, wrap "Inappropriate"), |
|
152 |
(OutOfResources, wrap "ResourceOut"), |
|
153 |
(OutOfResources, wrap "MemoryOut"), |
|
154 |
(Interrupted, wrap "Forced"), |
|
155 |
(Interrupted, wrap "User")] |
|
156 |
||
157 |
val known_szs_status_failures = known_szs_failures (prefix "SZS status ") |
|
158 |
val known_says_failures = known_szs_failures (prefix " says ") |
|
159 |
||
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
160 |
(* named ATPs *) |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
161 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
162 |
val agsyholN = "agsyhol" |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
163 |
val alt_ergoN = "alt_ergo" |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
164 |
val dummy_thfN = "dummy_thf" (* for experiments *) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
165 |
val eN = "e" |
48651 | 166 |
val e_malesN = "e_males" |
50927 | 167 |
val e_parN = "e_par" |
44590 | 168 |
val e_sineN = "e_sine" |
169 |
val e_tofofN = "e_tofof" |
|
45338 | 170 |
val iproverN = "iprover" |
171 |
val iprover_eqN = "iprover_eq" |
|
44099 | 172 |
val leo2N = "leo2" |
173 |
val satallaxN = "satallax" |
|
44590 | 174 |
val snarkN = "snark" |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
175 |
val spassN = "spass" |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
176 |
val spass_pirateN = "spass_pirate" |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
177 |
val vampireN = "vampire" |
44590 | 178 |
val waldmeisterN = "waldmeister" |
44423
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents:
44422
diff
changeset
|
179 |
val z3_tptpN = "z3_tptp" |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
180 |
val remote_prefix = "remote_" |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
181 |
|
38023 | 182 |
structure Data = Theory_Data |
183 |
( |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
184 |
type T = ((unit -> atp_config) * stamp) Symtab.table |
38023 | 185 |
val empty = Symtab.empty |
186 |
val extend = I |
|
46407
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents:
46402
diff
changeset
|
187 |
fun merge data : T = |
30e9720cc0b9
optimization: slice caching in case two consecutive slices are nearly identical
blanchet
parents:
46402
diff
changeset
|
188 |
Symtab.merge (eq_snd (op =)) data |
38023 | 189 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") |
190 |
) |
|
38017
3ad3e3ca2451
move Sledgehammer-specific code out of "Sledgehammer_TPTP_Format"
blanchet
parents:
38015
diff
changeset
|
191 |
|
43981
404ae49ce29f
give E at least two seconds -- anything else risks causing too early timeouts in the minimizer, because of too conservative time computations in E and eproof scripts
blanchet
parents:
43850
diff
changeset
|
192 |
fun to_secs min time = Int.max (min, (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
|
193 |
|
43473
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
194 |
val sosN = "sos" |
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
195 |
val no_sosN = "no_sos" |
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
196 |
|
44099 | 197 |
val force_sos = Attrib.setup_config_bool @{binding atp_force_sos} (K false) |
198 |
||
47032 | 199 |
val smartN = "smart" |
47073
c73f7b0c7ebc
generate weights and precedences for predicates as well
blanchet
parents:
47055
diff
changeset
|
200 |
(* val kboN = "kbo" *) |
47032 | 201 |
val lpoN = "lpo" |
47034
77da780ddd6b
implement term order attribute (for experiments)
blanchet
parents:
47033
diff
changeset
|
202 |
val xweightsN = "_weights" |
77da780ddd6b
implement term order attribute (for experiments)
blanchet
parents:
47033
diff
changeset
|
203 |
val xprecN = "_prec" |
77da780ddd6b
implement term order attribute (for experiments)
blanchet
parents:
47033
diff
changeset
|
204 |
val xsimpN = "_simp" (* SPASS-specific *) |
47032 | 205 |
|
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
206 |
(* Possible values for "atp_term_order": |
47049 | 207 |
"smart", "(kbo|lpo)(_weights)?(_prec|_simp)?" *) |
47032 | 208 |
val term_order = |
209 |
Attrib.setup_config_string @{binding atp_term_order} (K smartN) |
|
210 |
||
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
211 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
212 |
(* agsyHOL *) |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
213 |
|
54197
994ebb795b75
use definitions for LEO-II as well -- this simplifies the code and matches some users' expectations
blanchet
parents:
53586
diff
changeset
|
214 |
val agsyhol_thf0 = THF (Monomorphic, THF_Without_Choice) |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
215 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
216 |
val agsyhol_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
217 |
{exec = K (["AGSYHOL_HOME"], ["agsyHOL"]), |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
218 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
219 |
"--proof --time-out " ^ string_of_int (to_secs 1 timeout) ^ " " ^ |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
220 |
file_name, |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
221 |
proof_delims = tstp_proof_delims, |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
222 |
known_failures = known_szs_status_failures, |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
223 |
prem_role = Hypothesis, |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
224 |
best_slices = |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
225 |
(* FUDGE *) |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
226 |
K [(1.0, (((60, ""), agsyhol_thf0, "mono_native_higher", keep_lamsN, false), ""))], |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
227 |
best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
53515
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents:
53225
diff
changeset
|
228 |
best_max_new_mono_instances = default_max_new_mono_instances} |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
229 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
230 |
val agsyhol = (agsyholN, fn () => agsyhol_config) |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
231 |
|
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
232 |
|
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
233 |
(* Alt-Ergo *) |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
234 |
|
52995
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
blanchet
parents:
52754
diff
changeset
|
235 |
val alt_ergo_tff1 = TFF Polymorphic |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
236 |
|
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
237 |
val alt_ergo_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
238 |
{exec = K (["WHY3_HOME"], ["why3"]), |
50927 | 239 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
56379
d8ecce5d51cd
use Alt-Ergo 0.95.2, the latest and greatest version
blanchet
parents:
56378
diff
changeset
|
240 |
"--format tptp --prover 'Alt-Ergo,0.95.2,' --timelimit " ^ |
50927 | 241 |
string_of_int (to_secs 1 timeout) ^ " " ^ file_name, |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
242 |
proof_delims = [], |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
243 |
known_failures = |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
244 |
[(ProofMissing, ": Valid"), |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
245 |
(TimedOut, ": Timeout"), |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
246 |
(GaveUp, ": Unknown")], |
47976 | 247 |
prem_role = Hypothesis, |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
248 |
best_slices = fn _ => |
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
249 |
(* FUDGE *) |
51011 | 250 |
[(1.0, (((100, ""), alt_ergo_tff1, "poly_native", liftingN, false), ""))], |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
251 |
best_max_mono_iters = default_max_mono_iters, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
252 |
best_max_new_mono_instances = default_max_new_mono_instances} |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
253 |
|
47646 | 254 |
val alt_ergo = (alt_ergoN, fn () => alt_ergo_config) |
46643
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
255 |
|
a88bccd2b567
added support for Alt-Ergo through Why3 (mostly for experimental purposes, e.g. polymorphism vs. monomorphization)
blanchet
parents:
46481
diff
changeset
|
256 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
257 |
(* E *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
258 |
|
51873 | 259 |
fun is_e_at_least_1_3 () = string_ord (getenv "E_VERSION", "1.3") <> LESS |
260 |
fun is_e_at_least_1_6 () = string_ord (getenv "E_VERSION", "1.6") <> LESS |
|
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
261 |
fun is_e_at_least_1_8 () = string_ord (getenv "E_VERSION", "1.8") <> LESS |
44420 | 262 |
|
43473
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
263 |
val e_smartN = "smart" |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
264 |
val e_autoN = "auto" |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
265 |
val e_fun_weightN = "fun_weight" |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
266 |
val e_sym_offset_weightN = "sym_offset_weight" |
41725
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents:
41335
diff
changeset
|
267 |
|
47032 | 268 |
val e_selection_heuristic = |
269 |
Attrib.setup_config_string @{binding atp_e_selection_heuristic} (K e_smartN) |
|
41770 | 270 |
(* FUDGE *) |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
271 |
val e_default_fun_weight = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
272 |
Attrib.setup_config_real @{binding atp_e_default_fun_weight} (K 20.0) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
273 |
val e_fun_weight_base = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
274 |
Attrib.setup_config_real @{binding atp_e_fun_weight_base} (K 0.0) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
275 |
val e_fun_weight_span = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
276 |
Attrib.setup_config_real @{binding atp_e_fun_weight_span} (K 40.0) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
277 |
val e_default_sym_offs_weight = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
278 |
Attrib.setup_config_real @{binding atp_e_default_sym_offs_weight} (K 1.0) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
279 |
val e_sym_offs_weight_base = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
280 |
Attrib.setup_config_real @{binding atp_e_sym_offs_weight_base} (K ~20.0) |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
281 |
val e_sym_offs_weight_span = |
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
282 |
Attrib.setup_config_real @{binding atp_e_sym_offs_weight_span} (K 60.0) |
41725
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents:
41335
diff
changeset
|
283 |
|
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
284 |
fun e_selection_heuristic_case heuristic fw sow = |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
285 |
if heuristic = e_fun_weightN then fw |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
286 |
else if heuristic = e_sym_offset_weightN then sow |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
287 |
else raise Fail ("unexpected " ^ quote heuristic) |
41725
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents:
41335
diff
changeset
|
288 |
|
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
289 |
fun scaled_e_selection_weight ctxt heuristic w = |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
290 |
w * Config.get ctxt (e_selection_heuristic_case heuristic |
47029 | 291 |
e_fun_weight_span e_sym_offs_weight_span) |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
292 |
+ Config.get ctxt (e_selection_heuristic_case heuristic |
47029 | 293 |
e_fun_weight_base e_sym_offs_weight_base) |
41725
7cca2de89296
added support for bleeding-edge E weighting function "SymOffsetsWeight"
blanchet
parents:
41335
diff
changeset
|
294 |
|> Real.ceil |> signed_string_of_int |
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41269
diff
changeset
|
295 |
|
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
296 |
fun e_selection_weight_arguments ctxt heuristic sel_weights = |
51631 | 297 |
if heuristic = e_fun_weightN orelse heuristic = e_sym_offset_weightN then |
43622 | 298 |
(* supplied by Stephan Schulz *) |
41314
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
blanchet
parents:
41313
diff
changeset
|
299 |
"--split-clauses=4 --split-reuse-defs --simul-paramod --forward-context-sr \ |
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
blanchet
parents:
41313
diff
changeset
|
300 |
\--destructive-er-aggressive --destructive-er --presat-simplify \ |
47505
e33d957ae2bf
avoid option introduced in E 1.2 when invoking older versions of E
blanchet
parents:
47499
diff
changeset
|
301 |
\--prefer-initial-clauses -winvfreqrank -c1 -Ginvfreqconjmax -F1 \ |
e33d957ae2bf
avoid option introduced in E 1.2 when invoking older versions of E
blanchet
parents:
47499
diff
changeset
|
302 |
\--delete-bad-limit=150000000 -WSelectMaxLComplexAvoidPosPred -H'(4*" ^ |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
303 |
e_selection_heuristic_case heuristic "FunWeight" "SymOffsetWeight" ^ |
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48232
diff
changeset
|
304 |
"(SimulateSOS," ^ |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
305 |
(e_selection_heuristic_case heuristic |
47029 | 306 |
e_default_fun_weight e_default_sym_offs_weight |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
307 |
|> Config.get ctxt |> Real.ceil |> signed_string_of_int) ^ |
41314
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
blanchet
parents:
41313
diff
changeset
|
308 |
",20,1.5,1.5,1" ^ |
47030 | 309 |
(sel_weights () |
47029 | 310 |
|> map (fn (s, w) => "," ^ s ^ ":" ^ |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
311 |
scaled_e_selection_weight ctxt heuristic w) |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
312 |
|> implode) ^ |
41314
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
blanchet
parents:
41313
diff
changeset
|
313 |
"),3*ConjectureGeneralSymbolWeight(PreferNonGoals,200,100,200,50,50,1,100,\ |
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
blanchet
parents:
41313
diff
changeset
|
314 |
\1.5,1.5,1),1*Clauseweight(PreferProcessed,1,1,1),1*\ |
2dc1dfc1bc69
use the options provided by Stephan Schulz -- much better
blanchet
parents:
41313
diff
changeset
|
315 |
\FIFOWeight(PreferProcessed))'" |
51631 | 316 |
else |
317 |
"-xAuto" |
|
41313
a96ac4d180b7
optionally supply constant weights to E -- turned off by default until properly parameterized
blanchet
parents:
41269
diff
changeset
|
318 |
|
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
319 |
val e_ord_weights = |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
320 |
map (fn (s, w) => s ^ ":" ^ string_of_int w) #> space_implode "," |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
321 |
fun e_ord_precedence [_] = "" |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
322 |
| e_ord_precedence info = info |> map fst |> space_implode "<" |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
323 |
|
47039
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
blanchet
parents:
47038
diff
changeset
|
324 |
fun e_term_order_info_arguments false false _ = "" |
1b36a05a070d
added "metis_advisory_simp" option to orient as many equations as possible in Metis the right way (cf. "More SPASS with Isabelle")
blanchet
parents:
47038
diff
changeset
|
325 |
| e_term_order_info_arguments gen_weights gen_prec ord_info = |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
326 |
let val ord_info = ord_info () in |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
327 |
(if gen_weights then "--order-weights='" ^ e_ord_weights ord_info ^ "' " |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
328 |
else "") ^ |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
329 |
(if gen_prec then "--precedence='" ^ e_ord_precedence ord_info ^ "' " |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
330 |
else "") |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
331 |
end |
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
332 |
|
47032 | 333 |
fun effective_e_selection_heuristic ctxt = |
48653 | 334 |
if is_e_at_least_1_3 () then Config.get ctxt e_selection_heuristic |
48232
712d49104b13
don't ask E to generate a detailed proofs if not needed
blanchet
parents:
48131
diff
changeset
|
335 |
else e_autoN |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42332
diff
changeset
|
336 |
|
48653 | 337 |
fun e_kbo () = if is_e_at_least_1_3 () then "KBO6" else "KBO" |
47505
e33d957ae2bf
avoid option introduced in E 1.2 when invoking older versions of E
blanchet
parents:
47499
diff
changeset
|
338 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
339 |
val e_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
340 |
{exec = (fn () => (["E_HOME"], |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
341 |
if is_e_at_least_1_8 () then ["eprover"] else ["eproof_ram", "eproof"])), |
48651 | 342 |
arguments = fn ctxt => fn full_proof => fn heuristic => fn timeout => |
50927 | 343 |
fn file_name => |
48651 | 344 |
fn ({is_lpo, gen_weights, gen_prec, ...}, ord_info, sel_weights) => |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
345 |
"--tstp-in --tstp-out --silent " ^ |
48651 | 346 |
e_selection_weight_arguments ctxt heuristic sel_weights ^ " " ^ |
347 |
e_term_order_info_arguments gen_weights gen_prec ord_info ^ " " ^ |
|
348 |
"--term-ordering=" ^ (if is_lpo then "LPO4" else e_kbo ()) ^ " " ^ |
|
349 |
"--cpu-limit=" ^ string_of_int (to_secs 2 timeout) ^ |
|
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
350 |
(if is_e_at_least_1_8 () then |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
351 |
" --proof-object=1" |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
352 |
else |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
353 |
" --output-level=5" ^ |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
354 |
(if is_e_at_least_1_6 () then |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
355 |
" --pcl-shell-level=" ^ (if full_proof then "0" else "2") |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
356 |
else |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
357 |
"")) ^ |
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
358 |
" " ^ file_name, |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
359 |
proof_delims = |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
360 |
[("# SZS output start CNFRefutation", "# SZS output end CNFRefutation")] @ |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
361 |
tstp_proof_delims, |
36265
41c9e755e552
distinguish between the different ATP errors in the user interface;
blanchet
parents:
36264
diff
changeset
|
362 |
known_failures = |
45203 | 363 |
[(TimedOut, "Failure: Resource limit exceeded (time)"), |
47972 | 364 |
(TimedOut, "time limit exceeded")] @ |
365 |
known_szs_status_failures, |
|
47976 | 366 |
prem_role = Conjecture, |
42646
4781fcd53572
replaced some Unsynchronized.refs with Config.Ts
blanchet
parents:
42643
diff
changeset
|
367 |
best_slices = fn ctxt => |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
368 |
let val heuristic = effective_e_selection_heuristic ctxt in |
43474
423cd1ecf714
optimized E's time slicing, based on latest exhaustive Judgment Day results
blanchet
parents:
43473
diff
changeset
|
369 |
(* FUDGE *) |
47038
2409b484e1cc
continued implementation of term ordering attributes
blanchet
parents:
47034
diff
changeset
|
370 |
if heuristic = e_smartN then |
51017 | 371 |
[(0.15, (((128, meshN), FOF, "mono_tags??", combsN, false), e_fun_weightN)), |
372 |
(0.15, (((128, mashN), FOF, "mono_guards??", combsN, false), e_sym_offset_weightN)), |
|
373 |
(0.15, (((91, mepoN), FOF, "mono_tags??", combsN, false), e_autoN)), |
|
374 |
(0.15, (((1000, meshN), FOF, "poly_guards??", combsN, false), e_sym_offset_weightN)), |
|
51214
4fb12e2598dc
swap slices so that the last slice is more complete (for minimization)
blanchet
parents:
51205
diff
changeset
|
375 |
(0.15, (((256, mepoN), FOF, "mono_tags??", liftingN, false), e_fun_weightN)), |
4fb12e2598dc
swap slices so that the last slice is more complete (for minimization)
blanchet
parents:
51205
diff
changeset
|
376 |
(0.25, (((64, mashN), FOF, "mono_guards??", combsN, false), e_fun_weightN))] |
43473
fb2713b803e6
deal with ATP time slices in a more flexible/robust fashion
blanchet
parents:
43467
diff
changeset
|
377 |
else |
51011 | 378 |
[(1.0, (((500, ""), FOF, "mono_tags??", combsN, false), heuristic))] |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
379 |
end, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
380 |
best_max_mono_iters = default_max_mono_iters, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
381 |
best_max_new_mono_instances = default_max_new_mono_instances} |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
382 |
|
47646 | 383 |
val e = (eN, fn () => e_config) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
384 |
|
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
385 |
|
48651 | 386 |
(* E-MaLeS *) |
387 |
||
388 |
val e_males_config : atp_config = |
|
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
389 |
{exec = K (["E_MALES_HOME"], ["emales.py"]), |
50927 | 390 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
391 |
"-t " ^ string_of_int (to_secs 1 timeout) ^ " -p " ^ file_name, |
|
48651 | 392 |
proof_delims = tstp_proof_delims, |
393 |
known_failures = #known_failures e_config, |
|
394 |
prem_role = Conjecture, |
|
395 |
best_slices = |
|
396 |
(* FUDGE *) |
|
51018 | 397 |
K [(0.25, (((500, meshN), FOF, "mono_guards??", combs_or_liftingN, false), "")), |
398 |
(0.25, (((150, meshN), FOF, "poly_tags??", combs_or_liftingN, false), "")), |
|
399 |
(0.25, (((50, meshN), FOF, "mono_tags??", combs_or_liftingN, false), "")), |
|
400 |
(0.25, (((1000, meshN), FOF, "poly_guards??", combsN, false), ""))], |
|
48651 | 401 |
best_max_mono_iters = default_max_mono_iters, |
402 |
best_max_new_mono_instances = default_max_new_mono_instances} |
|
403 |
||
404 |
val e_males = (e_malesN, fn () => e_males_config) |
|
405 |
||
406 |
||
50927 | 407 |
(* E-Par *) |
408 |
||
409 |
val e_par_config : atp_config = |
|
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
410 |
{exec = K (["E_HOME"], ["runepar.pl"]), |
50927 | 411 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
412 |
string_of_int (to_secs 1 timeout) ^ " 1 " (* SInE *) ^ file_name ^ |
|
413 |
" 2" (* proofs *), |
|
414 |
proof_delims = tstp_proof_delims, |
|
415 |
known_failures = #known_failures e_config, |
|
416 |
prem_role = Conjecture, |
|
417 |
best_slices = #best_slices e_males_config, |
|
418 |
best_max_mono_iters = default_max_mono_iters, |
|
419 |
best_max_new_mono_instances = default_max_new_mono_instances} |
|
420 |
||
421 |
val e_par = (e_parN, fn () => e_par_config) |
|
422 |
||
423 |
||
48700 | 424 |
(* iProver *) |
425 |
||
426 |
val iprover_config : atp_config = |
|
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
427 |
{exec = K (["IPROVER_HOME"], ["iprover"]), |
50927 | 428 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
48720 | 429 |
"--clausifier \"$IPROVER_HOME\"/vclausify_rel --time_out_real " ^ |
50927 | 430 |
string_of_real (Time.toReal timeout) ^ " " ^ file_name, |
48700 | 431 |
proof_delims = tstp_proof_delims, |
432 |
known_failures = |
|
433 |
[(ProofIncomplete, "% SZS output start CNFRefutation")] @ |
|
434 |
known_szs_status_failures, |
|
435 |
prem_role = Hypothesis, |
|
436 |
best_slices = |
|
437 |
(* FUDGE *) |
|
51011 | 438 |
K [(1.0, (((150, ""), FOF, "mono_guards??", liftingN, false), ""))], |
48700 | 439 |
best_max_mono_iters = default_max_mono_iters, |
440 |
best_max_new_mono_instances = default_max_new_mono_instances} |
|
441 |
||
442 |
val iprover = (iproverN, fn () => iprover_config) |
|
443 |
||
444 |
||
445 |
(* iProver-Eq *) |
|
446 |
||
447 |
val iprover_eq_config : atp_config = |
|
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
448 |
{exec = K (["IPROVER_EQ_HOME"], ["iprover-eq"]), |
48700 | 449 |
arguments = #arguments iprover_config, |
450 |
proof_delims = #proof_delims iprover_config, |
|
451 |
known_failures = #known_failures iprover_config, |
|
452 |
prem_role = #prem_role iprover_config, |
|
453 |
best_slices = #best_slices iprover_config, |
|
454 |
best_max_mono_iters = #best_max_mono_iters iprover_config, |
|
455 |
best_max_new_mono_instances = #best_max_new_mono_instances iprover_config} |
|
456 |
||
457 |
val iprover_eq = (iprover_eqN, fn () => iprover_eq_config) |
|
458 |
||
459 |
||
44099 | 460 |
(* LEO-II *) |
461 |
||
48004
989a34fa72b3
don't generate definitions for LEO-II -- this cuases more harm than good
blanchet
parents:
47985
diff
changeset
|
462 |
(* LEO-II supports definitions, but it performs significantly better on our |
989a34fa72b3
don't generate definitions for LEO-II -- this cuases more harm than good
blanchet
parents:
47985
diff
changeset
|
463 |
benchmarks when they are not used. *) |
54197
994ebb795b75
use definitions for LEO-II as well -- this simplifies the code and matches some users' expectations
blanchet
parents:
53586
diff
changeset
|
464 |
val leo2_thf0 = THF (Monomorphic, THF_Without_Choice) |
44754 | 465 |
|
44099 | 466 |
val leo2_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
467 |
{exec = K (["LEO2_HOME"], ["leo.opt", "leo"]), |
50927 | 468 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
48651 | 469 |
"--foatp e --atp e=\"$E_HOME\"/eprover \ |
470 |
\--atp epclextract=\"$E_HOME\"/epclextract \ |
|
50927 | 471 |
\--proofoutput 1 --timeout " ^ string_of_int (to_secs 1 timeout) ^ " " ^ |
472 |
file_name, |
|
44099 | 473 |
proof_delims = tstp_proof_delims, |
45207 | 474 |
known_failures = |
47974
08d2dcc2dab9
improved LEO-II definition handling -- still hoping for a fix directly in LEO-II
blanchet
parents:
47972
diff
changeset
|
475 |
[(TimedOut, "CPU time limit exceeded, terminating"), |
47972 | 476 |
(GaveUp, "No.of.Axioms")] @ |
477 |
known_szs_status_failures, |
|
47976 | 478 |
prem_role = Hypothesis, |
47914
94f37848b7c9
LEO-II's "--sos" option confusingly disables rather than enables SOS, and SOS seems to be ignored anyway; also, pass a number of facts that's more appropriate for each prover
blanchet
parents:
47912
diff
changeset
|
479 |
best_slices = |
44099 | 480 |
(* FUDGE *) |
51011 | 481 |
K [(1.0, (((40, ""), leo2_thf0, "mono_native_higher", keep_lamsN, false), ""))], |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
482 |
best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
53515
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents:
53225
diff
changeset
|
483 |
best_max_new_mono_instances = default_max_new_mono_instances} |
39491
2416666e6f94
refactoring: move ATP proof and error extraction code to "ATP_Proof" module
blanchet
parents:
39375
diff
changeset
|
484 |
|
47646 | 485 |
val leo2 = (leo2N, fn () => leo2_config) |
44099 | 486 |
|
487 |
||
488 |
(* Satallax *) |
|
489 |
||
52097 | 490 |
(* Choice is disabled until there is proper reconstruction for it. *) |
54197
994ebb795b75
use definitions for LEO-II as well -- this simplifies the code and matches some users' expectations
blanchet
parents:
53586
diff
changeset
|
491 |
val satallax_thf0 = THF (Monomorphic, THF_Without_Choice) |
44754 | 492 |
|
44099 | 493 |
val satallax_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
494 |
{exec = K (["SATALLAX_HOME"], ["satallax.opt", "satallax"]), |
50927 | 495 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
496 |
"-p hocore -t " ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name, |
|
45162 | 497 |
proof_delims = |
498 |
[("% Higher-Order Unsat Core BEGIN", "% Higher-Order Unsat Core END")], |
|
45203 | 499 |
known_failures = known_szs_status_failures, |
47981 | 500 |
prem_role = Hypothesis, |
44416
cabd06b69c18
added formats to the slice and use TFF for remote Vampire
blanchet
parents:
44391
diff
changeset
|
501 |
best_slices = |
44754 | 502 |
(* FUDGE *) |
51011 | 503 |
K [(1.0, (((60, ""), satallax_thf0, "mono_native_higher", keep_lamsN, false), ""))], |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
504 |
best_max_mono_iters = default_max_mono_iters - 1 (* FUDGE *), |
53515
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents:
53225
diff
changeset
|
505 |
best_max_new_mono_instances = default_max_new_mono_instances} |
44099 | 506 |
|
47646 | 507 |
val satallax = (satallaxN, fn () => satallax_config) |
44099 | 508 |
|
509 |
||
510 |
(* SPASS *) |
|
42725
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents:
42723
diff
changeset
|
511 |
|
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
512 |
val spass_H1SOS = "-Heuristic=1 -SOS" |
50333
20c69b00e73c
tweak SPASS default a tiny bit, so that a more interesting heuristic is chosen when "slicing=false" (for experiments)
blanchet
parents:
49991
diff
changeset
|
513 |
val spass_H2 = "-Heuristic=2" |
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
514 |
val spass_H2LR0LT0 = "-Heuristic=2 -LR=0 -LT=0" |
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
515 |
val spass_H2NuVS0 = "-Heuristic=2 -RNuV=1 -Sorts=0" |
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
516 |
val spass_H2NuVS0Red2 = "-Heuristic=2 -RNuV=1 -Sorts=0 -RFRew=2 -RBRew=2 -RTaut=2" |
50333
20c69b00e73c
tweak SPASS default a tiny bit, so that a more interesting heuristic is chosen when "slicing=false" (for experiments)
blanchet
parents:
49991
diff
changeset
|
517 |
val spass_H2SOS = "-Heuristic=2 -SOS" |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
518 |
|
50950 | 519 |
val spass_extra_options = |
520 |
Attrib.setup_config_string @{binding atp_spass_extra_options} (K "") |
|
521 |
||
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
522 |
(* FIXME: Make "SPASS_NEW_HOME" legacy. *) |
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
523 |
val spass_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
524 |
{exec = K (["SPASS_NEW_HOME", "SPASS_HOME"], ["SPASS"]), |
52151
de43876e77bf
disable SPASS's splitting if Isar proofs are desired, because these are not handled by the proof reconstruction code (and it's not clear how to handle them considering the lack of documentation)
blanchet
parents:
52097
diff
changeset
|
525 |
arguments = fn _ => fn full_proofs => fn extra_options => fn timeout => |
de43876e77bf
disable SPASS's splitting if Isar proofs are desired, because these are not handled by the proof reconstruction code (and it's not clear how to handle them considering the lack of documentation)
blanchet
parents:
52097
diff
changeset
|
526 |
fn file_name => fn _ => |
de43876e77bf
disable SPASS's splitting if Isar proofs are desired, because these are not handled by the proof reconstruction code (and it's not clear how to handle them considering the lack of documentation)
blanchet
parents:
52097
diff
changeset
|
527 |
"-Isabelle=1 " ^ (if full_proofs then "-Splits=0 " else "") ^ |
de43876e77bf
disable SPASS's splitting if Isar proofs are desired, because these are not handled by the proof reconstruction code (and it's not clear how to handle them considering the lack of documentation)
blanchet
parents:
52097
diff
changeset
|
528 |
"-TimeLimit=" ^ string_of_int (to_secs 1 timeout) ^ " " ^ file_name |
de43876e77bf
disable SPASS's splitting if Isar proofs are desired, because these are not handled by the proof reconstruction code (and it's not clear how to handle them considering the lack of documentation)
blanchet
parents:
52097
diff
changeset
|
529 |
|> extra_options <> "" ? prefix (extra_options ^ " "), |
36369
d2cd0d04b8e6
handle ATP proof delimiters in a cleaner, more extensible fashion
blanchet
parents:
36289
diff
changeset
|
530 |
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
|
531 |
known_failures = |
53225
16235bb41881
got rid of old error -- users who install SPASS manually are responsible for any version mismatches
blanchet
parents:
52995
diff
changeset
|
532 |
[(GaveUp, "SPASS beiseite: Completion found"), |
36370
a4f601daa175
centralized ATP-specific error handling in "atp_wrapper.ML"
blanchet
parents:
36369
diff
changeset
|
533 |
(TimedOut, "SPASS beiseite: Ran out of time"), |
36965 | 534 |
(OutOfResources, "SPASS beiseite: Maximal number of loops exceeded"), |
37413 | 535 |
(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
|
536 |
(MalformedInput, "Free Variable"), |
44391 | 537 |
(Unprovable, "No formulae and clauses found in input file"), |
47972 | 538 |
(InternalError, "Please report this error")] @ |
539 |
known_perl_failures, |
|
47976 | 540 |
prem_role = Conjecture, |
50950 | 541 |
best_slices = fn ctxt => |
42723 | 542 |
(* FUDGE *) |
51016 | 543 |
[(0.1667, (((150, meshN), DFG Monomorphic, "mono_native", combsN, true), "")), |
544 |
(0.1667, (((500, meshN), DFG Monomorphic, "mono_native", liftingN, true), spass_H2SOS)), |
|
545 |
(0.1666, (((50, meshN), DFG Monomorphic, "mono_native", liftingN, true), spass_H2LR0LT0)), |
|
546 |
(0.1000, (((250, meshN), DFG Monomorphic, "mono_native", combsN, true), spass_H2NuVS0)), |
|
547 |
(0.1000, (((1000, mepoN), DFG Monomorphic, "mono_native", liftingN, true), spass_H1SOS)), |
|
548 |
(0.1000, (((150, meshN), DFG Monomorphic, "poly_guards??", liftingN, false), spass_H2NuVS0Red2)), |
|
549 |
(0.1000, (((300, meshN), DFG Monomorphic, "mono_native", combsN, true), spass_H2SOS)), |
|
550 |
(0.1000, (((100, meshN), DFG Monomorphic, "mono_native", combs_and_liftingN, true), spass_H2))] |
|
50950 | 551 |
|> (case Config.get ctxt spass_extra_options of |
552 |
"" => I |
|
553 |
| opts => map (apsnd (apsnd (K opts)))), |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
554 |
best_max_mono_iters = default_max_mono_iters, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
555 |
best_max_new_mono_instances = default_max_new_mono_instances} |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
556 |
|
48005
eeede26f2721
killed SPASS 3.5/3.7 FLOTTER hack -- requires users to upgrade to SPASS 3.8
blanchet
parents:
48004
diff
changeset
|
557 |
val spass = (spassN, fn () => spass_config) |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
558 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
559 |
|
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
560 |
(* Vampire *) |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
561 |
|
56380
9bb2856cc845
don't pass Vampire option that doesn't exist anymore (and that wasn't strictly necessary with older Vampires)
blanchet
parents:
56379
diff
changeset
|
562 |
(* Vampire 1.8 has TFF0 support, but the support was buggy until revision |
48007 | 563 |
1435 (or shortly before). *) |
51873 | 564 |
fun is_vampire_at_least_1_8 () = string_ord (getenv "VAMPIRE_VERSION", "1.8") <> LESS |
565 |
fun is_vampire_beyond_1_8 () = string_ord (getenv "VAMPIRE_VERSION", "1.8") = GREATER |
|
44420 | 566 |
|
52995
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
blanchet
parents:
52754
diff
changeset
|
567 |
val vampire_tff0 = TFF Monomorphic |
44589
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents:
44586
diff
changeset
|
568 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
569 |
val vampire_config : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
570 |
{exec = K (["VAMPIRE_HOME"], ["vampire"]), |
51872
af2e0b2c4d7e
pass certain readability-enhancing Vampire options only when an Isar proof is needed
blanchet
parents:
51631
diff
changeset
|
571 |
arguments = fn _ => fn full_proof => fn sos => fn timeout => fn file_name => |
af2e0b2c4d7e
pass certain readability-enhancing Vampire options only when an Isar proof is needed
blanchet
parents:
51631
diff
changeset
|
572 |
fn _ => |
48651 | 573 |
"--mode casc -t " ^ string_of_int (to_secs 1 timeout) ^ |
48653 | 574 |
" --proof tptp --output_axiom_names on" ^ |
575 |
(if is_vampire_at_least_1_8 () then |
|
51196
1ff19dfd3111
turn off more evil Vampire options to facilitate Isar proof generation
blanchet
parents:
51018
diff
changeset
|
576 |
(* Cf. p. 20 of http://www.complang.tuwien.ac.at/lkovacs/Cade23_Tutorial_Slides/Session2_Slides.pdf *) |
51872
af2e0b2c4d7e
pass certain readability-enhancing Vampire options only when an Isar proof is needed
blanchet
parents:
51631
diff
changeset
|
577 |
(if full_proof then |
56848
67e6803e3765
added missing space between command-line options
blanchet
parents:
56404
diff
changeset
|
578 |
" --forced_options splitting=off:equality_proxy=off:general_splitting=off\ |
51872
af2e0b2c4d7e
pass certain readability-enhancing Vampire options only when an Isar proof is needed
blanchet
parents:
51631
diff
changeset
|
579 |
\:inequality_splitting=0:naming=0" |
af2e0b2c4d7e
pass certain readability-enhancing Vampire options only when an Isar proof is needed
blanchet
parents:
51631
diff
changeset
|
580 |
else |
af2e0b2c4d7e
pass certain readability-enhancing Vampire options only when an Isar proof is needed
blanchet
parents:
51631
diff
changeset
|
581 |
"") |
48653 | 582 |
else |
583 |
"") ^ |
|
50927 | 584 |
" --thanks \"Andrei and Krystof\" --input_file " ^ file_name |
48651 | 585 |
|> sos = sosN ? prefix "--sos on ", |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
586 |
proof_delims = |
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
587 |
[("=========== Refutation ==========", |
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
588 |
"======= End of refutation =======")] @ |
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
589 |
tstp_proof_delims, |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
590 |
known_failures = |
43050
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
42999
diff
changeset
|
591 |
[(GaveUp, "UNPROVABLE"), |
59284a13abc4
support "metis" and "metisFT" as provers in the architecture, so they can be used for minimizing
blanchet
parents:
42999
diff
changeset
|
592 |
(GaveUp, "CANNOT PROVE"), |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
593 |
(Unprovable, "Satisfiability detected"), |
38647
5500241da479
play with fudge factor + parse one more Vampire error
blanchet
parents:
38646
diff
changeset
|
594 |
(Unprovable, "Termination reason: Satisfiable"), |
47972 | 595 |
(Interrupted, "Aborted by signal SIGINT")] @ |
596 |
known_szs_status_failures, |
|
47976 | 597 |
prem_role = Conjecture, |
42725
64dea91bbe0e
added "force_sos" options to control SPASS's and Vampire's use of SOS in experiments + added corresponding Mirabelle options
blanchet
parents:
42723
diff
changeset
|
598 |
best_slices = fn ctxt => |
42723 | 599 |
(* FUDGE *) |
48653 | 600 |
(if is_vampire_beyond_1_8 () then |
51016 | 601 |
[(0.333, (((500, meshN), vampire_tff0, "mono_guards??", combs_or_liftingN, false), sosN)), |
602 |
(0.333, (((150, meshN), vampire_tff0, "poly_tags??", combs_or_liftingN, false), sosN)), |
|
51017 | 603 |
(0.334, (((50, meshN), vampire_tff0, "mono_native", combs_or_liftingN, false), no_sosN))] |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
604 |
else |
51016 | 605 |
[(0.333, (((150, meshN), FOF, "poly_guards??", combs_or_liftingN, false), sosN)), |
51011 | 606 |
(0.333, (((500, meshN), FOF, "mono_tags??", combs_or_liftingN, false), sosN)), |
51016 | 607 |
(0.334, (((50, meshN), FOF, "mono_guards??", combs_or_liftingN, false), no_sosN))]) |
44099 | 608 |
|> (if Config.get ctxt force_sos then hd #> apfst (K 1.0) #> single |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
609 |
else I), |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
610 |
best_max_mono_iters = default_max_mono_iters, |
53515
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents:
53225
diff
changeset
|
611 |
best_max_new_mono_instances = 2 * default_max_new_mono_instances (* FUDGE *)} |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
612 |
|
47646 | 613 |
val vampire = (vampireN, fn () => vampire_config) |
37509
f39464d971c4
factor out TPTP format output into file of its own, to facilitate further changes
blanchet
parents:
37506
diff
changeset
|
614 |
|
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
615 |
|
48803
ffa31bf5c662
tone down "z3_tptp", now that Z3 (starting with 4.1) no longer supports TPTP TFF0
blanchet
parents:
48801
diff
changeset
|
616 |
(* Z3 with TPTP syntax (half experimental, half legacy) *) |
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
617 |
|
52995
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
blanchet
parents:
52754
diff
changeset
|
618 |
val z3_tff0 = TFF Monomorphic |
44589
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents:
44586
diff
changeset
|
619 |
|
44423
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents:
44422
diff
changeset
|
620 |
val z3_tptp_config : atp_config = |
56378 | 621 |
{exec = K (["Z3_TPTP_HOME"], ["z3_tptp"]), |
50927 | 622 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
56404
9cb137ec6ec8
use Z3 TPTP cores rather than proofs since the latter are somewhat broken
blanchet
parents:
56397
diff
changeset
|
623 |
"-core -t:" ^ string_of_int (to_secs 1 timeout) ^ " -file:" ^ file_name, |
56397 | 624 |
proof_delims = [("SZS status Theorem", "")], |
45203 | 625 |
known_failures = known_szs_status_failures, |
47976 | 626 |
prem_role = Hypothesis, |
42723 | 627 |
best_slices = |
44423
f74707e12d30
exploit TFF format in Z3 used as ATP, and renamed it "z3_tptp"
blanchet
parents:
44422
diff
changeset
|
628 |
(* FUDGE *) |
51011 | 629 |
K [(0.5, (((250, meshN), z3_tff0, "mono_native", combsN, false), "")), |
630 |
(0.25, (((125, mepoN), z3_tff0, "mono_native", combsN, false), "")), |
|
631 |
(0.125, (((62, mashN), z3_tff0, "mono_native", combsN, false), "")), |
|
632 |
(0.125, (((31, meshN), z3_tff0, "mono_native", combsN, false), ""))], |
|
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
633 |
best_max_mono_iters = default_max_mono_iters, |
53515
f5b678b155f6
adjusted number of generated monomorphic instances for new monomorphizer based on new evaluation (E, SPASS, Vampire)
blanchet
parents:
53225
diff
changeset
|
634 |
best_max_new_mono_instances = 2 * default_max_new_mono_instances (* FUDGE *)} |
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
635 |
|
47646 | 636 |
val z3_tptp = (z3_tptpN, fn () => z3_tptp_config) |
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
637 |
|
44590 | 638 |
|
48131 | 639 |
(* Not really a prover: Experimental Polymorphic THF and DFG output *) |
44590 | 640 |
|
51919
097b191d1f0d
use right default for "uncurried_aliases" with polymorphic SPASS
blanchet
parents:
51873
diff
changeset
|
641 |
fun dummy_config prem_role format type_enc uncurried_aliases : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
642 |
{exec = K (["ISABELLE_ATP"], ["scripts/dummy_atp"]), |
50927 | 643 |
arguments = K (K (K (K (K (K ""))))), |
44590 | 644 |
proof_delims = [], |
45203 | 645 |
known_failures = known_szs_status_failures, |
51467 | 646 |
prem_role = prem_role, |
45521 | 647 |
best_slices = |
51011 | 648 |
K [(1.0, (((200, ""), format, type_enc, |
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48715
diff
changeset
|
649 |
if is_format_higher_order format then keep_lamsN |
51919
097b191d1f0d
use right default for "uncurried_aliases" with polymorphic SPASS
blanchet
parents:
51873
diff
changeset
|
650 |
else combsN, uncurried_aliases), ""))], |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
651 |
best_max_mono_iters = default_max_mono_iters, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
652 |
best_max_new_mono_instances = default_max_new_mono_instances} |
44590 | 653 |
|
54197
994ebb795b75
use definitions for LEO-II as well -- this simplifies the code and matches some users' expectations
blanchet
parents:
53586
diff
changeset
|
654 |
val dummy_thf_format = THF (Polymorphic, THF_With_Choice) |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
655 |
val dummy_thf_config = dummy_config Hypothesis dummy_thf_format "poly_native_higher" false |
47646 | 656 |
val dummy_thf = (dummy_thfN, fn () => dummy_thf_config) |
44754 | 657 |
|
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
658 |
val spass_pirate_format = DFG Polymorphic |
54802 | 659 |
val remote_spass_pirate_config : atp_config = |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
660 |
{exec = K (["ISABELLE_ATP"], ["scripts/remote_spass_pirate"]), |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
661 |
arguments = fn _ => fn _ => fn _ => fn timeout => fn file_name => fn _ => |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
662 |
string_of_int (to_secs 1 timeout) ^ " " ^ file_name, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
663 |
proof_delims = [("Involved clauses:", "Involved clauses:")], |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
664 |
known_failures = known_szs_status_failures, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
665 |
prem_role = #prem_role spass_config, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
666 |
best_slices = K [(1.0, (((200, ""), spass_pirate_format, "tc_native", combsN, true), ""))], |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
667 |
best_max_mono_iters = default_max_mono_iters, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
668 |
best_max_new_mono_instances = default_max_new_mono_instances} |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
669 |
val remote_spass_pirate = (remote_prefix ^ spass_pirateN, fn () => remote_spass_pirate_config) |
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
670 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
671 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
672 |
(* Remote ATP invocation via SystemOnTPTP *) |
28596
fcd463a6b6de
tuned interfaces -- plain prover function, without thread;
wenzelm
parents:
28592
diff
changeset
|
673 |
|
49984 | 674 |
val remote_systems = Synchronized.var "atp_remote_systems" ([] : string list) |
31835 | 675 |
|
49984 | 676 |
fun get_remote_systems () = |
49987
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
677 |
TimeLimit.timeLimit (seconds 10.0) |
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
678 |
(fn () => |
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
679 |
case Isabelle_System.bash_output |
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
680 |
"\"$ISABELLE_ATP/scripts/remote_atp\" -w 2>&1" of |
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
681 |
(output, 0) => split_lines output |
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
682 |
| (output, _) => |
53586
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
683 |
(warning (case extract_known_atp_failure known_perl_failures output of |
bd5fa6425993
prefixed types and some functions with "atp_" for disambiguation
blanchet
parents:
53515
diff
changeset
|
684 |
SOME failure => string_of_atp_failure failure |
49987
25e333d2eccd
added a timeout around script that relies on the network
blanchet
parents:
49984
diff
changeset
|
685 |
| NONE => trim_line output ^ "."); [])) () |
49991
e0761153fbd1
less verbose -- the warning will reach the users anyway by other means
blanchet
parents:
49990
diff
changeset
|
686 |
handle TimeLimit.TimeOut => [] |
31835 | 687 |
|
49984 | 688 |
fun find_remote_system name [] systems = |
42537
25ceb855a18b
improve version handling -- prefer versions of ToFoF, SInE, and SNARK that are known to work
blanchet
parents:
42535
diff
changeset
|
689 |
find_first (String.isPrefix (name ^ "---")) systems |
49984 | 690 |
| find_remote_system name (version :: versions) systems = |
38690
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents:
38685
diff
changeset
|
691 |
case find_first (String.isPrefix (name ^ "---" ^ version)) systems of |
49984 | 692 |
NONE => find_remote_system name versions systems |
38690
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents:
38685
diff
changeset
|
693 |
| res => res |
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents:
38685
diff
changeset
|
694 |
|
49984 | 695 |
fun get_remote_system name versions = |
696 |
Synchronized.change_result remote_systems |
|
697 |
(fn systems => (if null systems then get_remote_systems () else systems) |
|
698 |
|> `(`(find_remote_system name versions))) |
|
32864
a226f29d4bdc
re-organized signature of AtpWrapper structure: records instead of unnamed parameters and return values,
boehmes
parents:
32740
diff
changeset
|
699 |
|
49984 | 700 |
fun the_remote_system name versions = |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
701 |
(case get_remote_system name versions of |
42955 | 702 |
(SOME sys, _) => sys |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
703 |
| (NONE, []) => error "SystemOnTPTP is not available." |
42955 | 704 |
| (NONE, syss) => |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
705 |
(case syss |> filter_out (String.isPrefix "%") |> filter_out (curry (op =) "") of |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
706 |
[] => error "SystemOnTPTP is currently not available." |
49990 | 707 |
| [msg] => error ("SystemOnTPTP is currently not available: " ^ msg ^ ".") |
46480 | 708 |
| syss => |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
709 |
error ("System " ^ quote name ^ " is not available at SystemOnTPTP.\n(Available systems: " ^ |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
710 |
commas_quote syss ^ ".)"))) |
31835 | 711 |
|
41148 | 712 |
val max_remote_secs = 240 (* give Geoff Sutcliffe's servers a break *) |
713 |
||
38690
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents:
38685
diff
changeset
|
714 |
fun remote_config system_name system_versions proof_delims known_failures |
47976 | 715 |
prem_role best_slice : atp_config = |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
716 |
{exec = K (["ISABELLE_ATP"], ["scripts/remote_atp"]), |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
717 |
arguments = fn _ => fn _ => fn command => fn timeout => fn file_name => fn _ => |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
718 |
(if command <> "" then "-c " ^ quote command ^ " " else "") ^ |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
719 |
"-s " ^ the_remote_system system_name system_versions ^ " " ^ |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
720 |
"-t " ^ string_of_int (Int.min (max_remote_secs, to_secs 1 timeout)) ^ |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
721 |
" " ^ file_name, |
42962
3b50fdeb6cfc
started adding support for THF output (but no lambdas)
blanchet
parents:
42955
diff
changeset
|
722 |
proof_delims = union (op =) tstp_proof_delims proof_delims, |
45203 | 723 |
known_failures = known_failures @ known_perl_failures @ known_says_failures, |
47976 | 724 |
prem_role = prem_role, |
48716
1d2a12bb0640
stop distinguishing between complete and incomplete slices, since this is very fragile and has hardly any useful semantics to users
blanchet
parents:
48715
diff
changeset
|
725 |
best_slices = fn ctxt => [(1.0, best_slice ctxt)], |
47962
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
726 |
best_max_mono_iters = default_max_mono_iters, |
137883567114
lower the monomorphization thresholds for less scalable provers
blanchet
parents:
47955
diff
changeset
|
727 |
best_max_new_mono_instances = default_max_new_mono_instances} |
42443
724e612ba248
implemented general slicing for ATPs, especially E 1.2w and above
blanchet
parents:
42332
diff
changeset
|
728 |
|
43500
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents:
43497
diff
changeset
|
729 |
fun remotify_config system_name system_versions best_slice |
47976 | 730 |
({proof_delims, known_failures, prem_role, ...} : atp_config) |
47912
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
blanchet
parents:
47900
diff
changeset
|
731 |
: atp_config = |
38690
38a926e033ad
make remote ATP versions more robust, by starting with "preferred" version numbers and falling back on any version
blanchet
parents:
38685
diff
changeset
|
732 |
remote_config system_name system_versions proof_delims known_failures |
47976 | 733 |
prem_role best_slice |
38023 | 734 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
735 |
fun remote_atp name system_name system_versions proof_delims known_failures |
47976 | 736 |
prem_role best_slice = |
40060
5ef6747aa619
first step in adding support for an SMT backend to Sledgehammer
blanchet
parents:
40059
diff
changeset
|
737 |
(remote_prefix ^ name, |
47912
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
blanchet
parents:
47900
diff
changeset
|
738 |
fn () => remote_config system_name system_versions proof_delims |
47976 | 739 |
known_failures prem_role best_slice) |
43500
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents:
43497
diff
changeset
|
740 |
fun remotify_atp (name, config) system_name system_versions best_slice = |
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents:
43497
diff
changeset
|
741 |
(remote_prefix ^ name, |
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
742 |
remotify_config system_name system_versions best_slice o config) |
28592 | 743 |
|
52995
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
blanchet
parents:
52754
diff
changeset
|
744 |
val explicit_tff0 = TFF Monomorphic |
44589
0a1dfc6365e9
first step towards polymorphic TFF + changed defaults for Vampire
blanchet
parents:
44586
diff
changeset
|
745 |
|
52094 | 746 |
val remote_agsyhol = |
747 |
remotify_atp agsyhol "agsyHOL" ["1.0", "1"] |
|
748 |
(K (((60, ""), agsyhol_thf0, "mono_native_higher", keep_lamsN, false), "") (* FUDGE *)) |
|
43500
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents:
43497
diff
changeset
|
749 |
val remote_e = |
52094 | 750 |
remotify_atp e "EP" ["1.7", "1.6", "1.5", "1"] |
51011 | 751 |
(K (((750, ""), FOF, "mono_tags??", combsN, false), "") (* FUDGE *)) |
48700 | 752 |
val remote_iprover = |
52094 | 753 |
remotify_atp iprover "iProver" ["0.99"] |
51011 | 754 |
(K (((150, ""), FOF, "mono_guards??", liftingN, false), "") (* FUDGE *)) |
48700 | 755 |
val remote_iprover_eq = |
52094 | 756 |
remotify_atp iprover_eq "iProver-Eq" ["0.8"] |
51011 | 757 |
(K (((150, ""), FOF, "mono_guards??", liftingN, false), "") (* FUDGE *)) |
44099 | 758 |
val remote_leo2 = |
52094 | 759 |
remotify_atp leo2 "LEO-II" ["1.5.0", "1.4", "1.3", "1.2", "1"] |
760 |
(K (((40, ""), leo2_thf0, "mono_native_higher", liftingN, false), "") (* FUDGE *)) |
|
44099 | 761 |
val remote_satallax = |
52094 | 762 |
remotify_atp satallax "Satallax" ["2.7", "2.3", "2"] |
763 |
(K (((60, ""), satallax_thf0, "mono_native_higher", keep_lamsN, false), "") (* FUDGE *)) |
|
43500
4c357b7aa710
provide appropriate type system and number of fact defaults for remote ATPs
blanchet
parents:
43497
diff
changeset
|
764 |
val remote_vampire = |
52995
ab98feb66684
Vampire 3.0 requires types to be declared -- make it happy (and get rid of "implicit" types since only Satallax seems to support them anymore)
blanchet
parents:
52754
diff
changeset
|
765 |
remotify_atp vampire "Vampire" ["3.0", "2.6", "2.5", "1.8"] |
51011 | 766 |
(K (((250, ""), vampire_tff0, "mono_native", combs_or_liftingN, false), "") (* FUDGE *)) |
44092
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
blanchet
parents:
43989
diff
changeset
|
767 |
val remote_e_sine = |
47912
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
blanchet
parents:
47900
diff
changeset
|
768 |
remote_atp e_sineN "SInE" ["0.4"] [] (#known_failures e_config) Conjecture |
51011 | 769 |
(K (((500, ""), FOF, "mono_guards??", combsN, false), "") (* FUDGE *)) |
41740
4b09f8b9e012
added "Z3 as an ATP" support to Sledgehammer locally
blanchet
parents:
41738
diff
changeset
|
770 |
val remote_snark = |
52094 | 771 |
remote_atp snarkN "SNARK" ["20120808r022", "20080805r029", "20080805r024"] |
47912
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
blanchet
parents:
47900
diff
changeset
|
772 |
[("refutation.", "end_refutation.")] [] Hypothesis |
51011 | 773 |
(K (((100, ""), explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *)) |
44092
bf489e54d7f8
renamed E wrappers for consistency with CASC conventions
blanchet
parents:
43989
diff
changeset
|
774 |
val remote_e_tofof = |
47912
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
blanchet
parents:
47900
diff
changeset
|
775 |
remote_atp e_tofofN "ToFoF" ["0.1"] [] (#known_failures e_config) Hypothesis |
51011 | 776 |
(K (((150, ""), explicit_tff0, "mono_native", liftingN, false), "") (* FUDGE *)) |
42938 | 777 |
val remote_waldmeister = |
778 |
remote_atp waldmeisterN "Waldmeister" ["710"] |
|
45521 | 779 |
[("#START OF PROOF", "Proved Goals:")] |
780 |
[(OutOfResources, "Too many function symbols"), |
|
47506 | 781 |
(Inappropriate, "**** Unexpected end of file."), |
45521 | 782 |
(Crashed, "Unrecoverable Segmentation Fault")] |
47912
12de57c5eee5
get rid of "conj_sym_kind" -- most interesting provers now have built-in sorts, and for the others (e.g. E) "Hypothesis" isn't too bad a default
blanchet
parents:
47900
diff
changeset
|
783 |
Hypothesis |
51011 | 784 |
(K (((50, ""), CNF_UEQ, "raw_mono_tags??", combsN, false), "") (* FUDGE *)) |
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
785 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
786 |
|
38454
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
787 |
(* Setup *) |
9043eefe8d71
detect old Vampire and give a nicer error message
blanchet
parents:
38433
diff
changeset
|
788 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
789 |
fun add_atp (name, config) thy = |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
790 |
Data.map (Symtab.update_new (name, (config, stamp ()))) thy |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
791 |
handle Symtab.DUP name => error ("Duplicate ATP: " ^ quote name ^ ".") |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
792 |
|
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
793 |
fun get_atp thy name = |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
794 |
fst (the (Symtab.lookup (Data.get thy) name)) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
795 |
handle Option.Option => error ("Unknown ATP: " ^ name ^ ".") |
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
796 |
|
41727
ab3f6d76fb23
available_provers ~> supported_provers (for clarity)
blanchet
parents:
41725
diff
changeset
|
797 |
val supported_atps = Symtab.keys o Data.get |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
798 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
799 |
fun is_atp_installed thy name = |
48376
416e4123baf3
use "eproof_ram" script if available (plug-in replacement for "eproof", but faster)
blanchet
parents:
48232
diff
changeset
|
800 |
let val {exec, ...} = get_atp thy name () in |
52754
d9d90d29860e
added support for E 1.8's internal proof objects (eliminating the need for "eproof_ram")
blanchet
parents:
52151
diff
changeset
|
801 |
exists (fn var => getenv var <> "") (fst (exec ())) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
802 |
end |
36371
8c83ea1a7740
move the Sledgehammer menu options to "sledgehammer_isar.ML"
blanchet
parents:
36370
diff
changeset
|
803 |
|
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
804 |
fun refresh_systems_on_tptp () = |
49984 | 805 |
Synchronized.change remote_systems (fn _ => get_remote_systems ()) |
40059
6ad9081665db
use consistent terminology in Sledgehammer: "prover = ATP or SMT solver or ..."
blanchet
parents:
39491
diff
changeset
|
806 |
|
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
807 |
fun effective_term_order ctxt atp = |
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
808 |
let val ord = Config.get ctxt term_order in |
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
809 |
if ord = smartN then |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
810 |
{is_lpo = false, gen_weights = (atp = spassN), gen_prec = (atp = spassN), |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
811 |
gen_simp = String.isSuffix spass_pirateN atp} |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
812 |
else |
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
813 |
let val is_lpo = String.isSubstring lpoN ord in |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
814 |
{is_lpo = is_lpo, gen_weights = not is_lpo andalso String.isSubstring xweightsN ord, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
815 |
gen_prec = String.isSubstring xprecN ord, gen_simp = String.isSubstring xsimpN ord} |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
816 |
end |
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
817 |
end |
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
818 |
|
52073
ccb292952774
started adding agsyHOL as an experimental prover
blanchet
parents:
51998
diff
changeset
|
819 |
val atps = |
54788
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
820 |
[agsyhol, alt_ergo, e, e_males, e_par, iprover, iprover_eq, leo2, satallax, spass, vampire, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
821 |
z3_tptp, dummy_thf, remote_agsyhol, remote_e, remote_e_sine, remote_e_tofof, remote_iprover, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
822 |
remote_iprover_eq, remote_leo2, remote_satallax, remote_vampire, remote_snark, |
a898e15b522a
primitive support for SPASS-Pirate (Daniel Wand's polymorphic SPASS prototype)
blanchet
parents:
54197
diff
changeset
|
823 |
remote_spass_pirate, remote_waldmeister] |
47055
16e2633f3b4b
made "spass" a "metaprover" that uses either the new SPASS or the old SPASS, to preserve backward compatibility and prepare for the upcoming release
blanchet
parents:
47053
diff
changeset
|
824 |
|
47606
06dde48a1503
true delayed evaluation of "SPASS_VERSION" environment variable
blanchet
parents:
47506
diff
changeset
|
825 |
val setup = fold add_atp atps |
35867 | 826 |
|
28592 | 827 |
end; |